├── Coccinelle ├── Licence.en.txt ├── Licence.fr.txt ├── basis │ ├── more_logic.v │ ├── terminaison.v │ ├── decidable_set.v │ └── ordered_set.v ├── examples │ └── cime_trace │ │ ├── arith_extension.v │ │ ├── closure_extension.v │ │ ├── rewriting.v │ │ ├── term_extension.v │ │ ├── make_or.ml │ │ └── more_list_extention.v ├── README ├── Makefile ├── INSTALL └── list_extensions │ └── equiv_list.v ├── Licence_CeCILL_V2.1-en.txt ├── Licence_CeCILL_V2.1-fr.txt ├── .github ├── dependabot.yml └── workflows │ └── main.yml ├── .gitignore ├── svn_mv_coq ├── doc ├── end.html └── begin.html ├── time_coqc ├── stat_time ├── Util ├── Logic │ ├── DepChoice.v │ ├── EpsilonUtil.v │ ├── IotaUtil.v │ ├── DepChoicePrf.v │ └── ClassicUtil.v ├── Relation │ ├── NotSN.v │ ├── RedLength.v │ ├── RelMidex.v │ ├── NotSN_IS.v │ ├── Cycle.v │ ├── SCC.v │ ├── RelSub.v │ ├── SCC_dec.v │ └── Preorder.v ├── List │ ├── ListSort.v │ └── ListMax.v ├── Pair │ └── PairUtil.v ├── Multiset │ ├── MultisetNat.v │ └── MultisetCore.v ├── Integer │ └── BigZUtil.v ├── Function │ ├── NaryFunction.v │ └── FunUtil.v ├── Nat │ ├── LeastNat.v │ ├── BigNUtil.v │ ├── Log2.v │ └── NatCompat.v ├── Vector │ ├── VecMax.v │ ├── VecBool.v │ └── VecFilterPerm.v ├── Option │ └── OptUtil.v └── Set_ │ └── InfSet.v ├── stat_time.awk ├── THANKS ├── Term ├── SimpleType │ ├── Terms.v │ ├── TermsSig.v │ ├── TermsEta.v │ └── TermsBuilding.v ├── Varyadic │ ├── VSubstitution.v │ ├── VTrs.v │ ├── VContext.v │ └── VSignature.v ├── String │ ├── SContext.v │ └── SReverse.v └── WithArity │ ├── ANotvar.v │ ├── ATrsNorm.v │ ├── AWFMInterpretation.v │ ├── ASubterm.v │ └── ASignature.v ├── ProofChecker ├── EmptyChecker.v ├── Problem.v ├── Proof.v ├── ProofChecker.v └── IntBasedChecker.v ├── stat_color ├── stat_coq ├── create_dist ├── .travis.yml ├── create_index ├── Makefile ├── COPYRIGHTS ├── NonTermin └── AVarCond.v ├── INSTALL ├── LICENSE ├── description ├── RPO ├── VPrecedence.v └── VLPO.v ├── SemLab └── ARootLab.v ├── opam ├── TODO ├── CONTRIBUTING ├── HORPO └── HorpoComp.v ├── PolyInt └── APolyInt_MA.v ├── SubtermCrit └── ASimpleProj.v └── MannaNess └── AMannaNess.v /Coccinelle/Licence.en.txt: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/color/master/Coccinelle/Licence.en.txt -------------------------------------------------------------------------------- /Coccinelle/Licence.fr.txt: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/color/master/Coccinelle/Licence.fr.txt -------------------------------------------------------------------------------- /Licence_CeCILL_V2.1-en.txt: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/color/master/Licence_CeCILL_V2.1-en.txt -------------------------------------------------------------------------------- /Licence_CeCILL_V2.1-fr.txt: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/fblanqui/color/master/Licence_CeCILL_V2.1-fr.txt -------------------------------------------------------------------------------- /.github/dependabot.yml: -------------------------------------------------------------------------------- 1 | version: 2 2 | 3 | updates: 4 | - package-ecosystem: github-actions 5 | directory: / 6 | schedule: 7 | interval: monthly 8 | -------------------------------------------------------------------------------- /Coccinelle/basis/more_logic.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Bool. 2 | 3 | 4 | Lemma impl_bool_def : 5 | forall p q, orb (negb p) q = true -> p = true -> q = true. 6 | Proof. 7 | intros p q H H0; subst;exact H. 8 | Qed. 9 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | doc/CoLoR.*.html 2 | doc/coqdoc.css 3 | doc/index.html 4 | doc/main.html 5 | .lia.cache 6 | .nia.cache 7 | .*.aux 8 | *.d 9 | *.glob 10 | *.vo 11 | CoqMakefile.conf 12 | rocq.mk 13 | rocq.mk.conf 14 | _CoqProject 15 | *.time 16 | *.vok 17 | *.vos 18 | make.log 19 | -------------------------------------------------------------------------------- /svn_mv_coq: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2015-03-27 7 | 8 | f=${1%.v} 9 | d=`dirname $2` 10 | g=`basename $2 .v` 11 | 12 | svn mv $f.v $d/$g.v 13 | for s in vo glob v.d time 14 | do 15 | mv $f.$s $d/$g.$s 16 | done 17 | -------------------------------------------------------------------------------- /doc/end.html: -------------------------------------------------------------------------------- 1 | 2 |

Index

3 | 4 | 5 |
6 | 7 |

8 | Valid XHTML 1.1

9 | 10 | 11 | 12 | 13 | -------------------------------------------------------------------------------- /time_coqc: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2015-03-19 7 | 8 | if test "$1" = "--print-version"; then coqc --print-version; exit 0; fi 9 | 10 | args=$* 11 | n=$# 12 | shift `expr $n - 1` 13 | f=${1%.v} 14 | /usr/bin/time --quiet -f "%e" -o $f.time coqc $args 15 | cat $f.time 16 | -------------------------------------------------------------------------------- /stat_time: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2015-03-19 7 | 8 | output=stat_time.log 9 | 10 | echo generate $output ... 11 | 12 | echo > $output 13 | 14 | for f in `find . -name \*.time` 15 | do 16 | echo `echo ${f%.time}.v` `cat $f` >> $output 17 | done 18 | 19 | awk -f stat_time.awk $output 20 | -------------------------------------------------------------------------------- /Util/Logic/DepChoice.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-08-08 6 | 7 | dependent choice axiom 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import RelUtil. 13 | 14 | Axiom dep_choice : forall (A : Type) (a : A) (R : relation A), 15 | classic_left_total R -> exists f, IS R f /\ f 0 = a. 16 | -------------------------------------------------------------------------------- /stat_time.awk: -------------------------------------------------------------------------------- 1 | # CoLoR, a Coq library on rewriting and termination. 2 | # See the COPYRIGHTS and LICENSE files. 3 | # 4 | # - Frederic Blanqui, 2015-03-19 5 | 6 | {if($2>10){subtotal+=$2;subnb++;print}} 7 | {total+=$2;nb++;if($2>maxval){maxval=$2;maxfile=$1}} 8 | END{printf"files: %d, time: %ds, mean: %.2fs\nmax: %.2fs (%s)\nfiles>10s: %d (%d%%), time: %ds (%d%%)\n",nb,total,total/nb,maxval,maxfile,subnb,100*subnb/nb,subtotal,100*subtotal/total} 9 | -------------------------------------------------------------------------------- /Util/Logic/EpsilonUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2014-12-11 6 | 7 | basic consequences of ClassicalEpsilon.constructive_indefinite_description 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import ClassicalEpsilon. 13 | 14 | Arguments constructive_indefinite_description [A P] _. 15 | 16 | Notation cid := constructive_indefinite_description. 17 | -------------------------------------------------------------------------------- /THANKS: -------------------------------------------------------------------------------- 1 | Frederic Blanqui thanks the Conseil Regional de Lorraine 2 | (http://www.cr-lorraine.fr/) for its grant in 2006, the Institute of 3 | Software of the Chinese Academy of Sciences 4 | (http://english.is.cas.cn/) for having hosted him between June 2012 5 | and August 2013, and the following people for their comments or help: 6 | Jean-Marc Notin, Loredana Tec, Paul Brauner, Eli Soubiran, Sorin 7 | Stratulat, Cédric Auger, Matthieu Sozeau, Anders Lundstedt, Pierre 8 | Ludmann, Vadim Zaliva. 9 | -------------------------------------------------------------------------------- /Term/SimpleType/Terms.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2006-04-27 6 | 7 | This file provides the development concerning terms of simply typed 8 | lambda-calculus. 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | From CoLoR Require TermsAlgebraic. 14 | 15 | Module Terms (Sig : TermsSig.Signature). 16 | 17 | Module Export TA := TermsAlgebraic.TermsAlgebraic Sig. 18 | 19 | End Terms. 20 | -------------------------------------------------------------------------------- /Coccinelle/examples/cime_trace/arith_extension.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith. 2 | 3 | Fixpoint le_bool (n m : nat ) : bool := 4 | match n,m with 5 | | 0,_ => true 6 | | _,0 => false 7 | | S n, S m => le_bool n m 8 | end. 9 | 10 | Lemma le_bool_ok_true : 11 | forall n m, le_bool n m = true -> le n m. 12 | Proof. 13 | induction n. 14 | 15 | intros m _;induction m. 16 | constructor. 17 | constructor;exact IHm. 18 | 19 | intros [|m]. 20 | intros abs;discriminate abs. 21 | simpl;intros h. 22 | apply le_n_S. 23 | apply IHn. exact h. 24 | Qed. 25 | -------------------------------------------------------------------------------- /.github/workflows/main.yml: -------------------------------------------------------------------------------- 1 | # following https://github.com/coq-community/docker-coq-action 2 | 3 | name: CI 4 | 5 | on: 6 | pull_request: 7 | types: [opened, synchronize, reopened] 8 | push: 9 | paths-ignore: 10 | - '**/*.md' 11 | workflow_dispatch: 12 | 13 | jobs: 14 | 15 | make: 16 | runs-on: ubuntu-latest 17 | strategy: 18 | fail-fast: false 19 | matrix: 20 | coq_version: [dev, 9.1, '9.0', '8.20', 8.19, 8.18, 8.17, 8.16] 21 | steps: 22 | - uses: actions/checkout@v6 23 | - uses: coq-community/docker-coq-action@v1 24 | with: 25 | opam_file: 'opam' 26 | coq_version: ${{ matrix.coq_version }} 27 | -------------------------------------------------------------------------------- /Util/Logic/IotaUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2014-12-11 6 | 7 | * Basic consequences of Description.constructive_definite_description 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import ClassicalDescription. 13 | 14 | Arguments constructive_definite_description [A P] _. 15 | 16 | Notation cdd := constructive_definite_description. 17 | Notation dec := excluded_middle_informative. 18 | 19 | Lemma dec1 A (P : A->Prop) x : {P x}+{~P x}. 20 | 21 | Proof. apply dec. Qed. 22 | 23 | Lemma eq_dec {A} (x y : A) : {x=y}+{x<>y}. 24 | 25 | Proof. apply dec. Qed. 26 | -------------------------------------------------------------------------------- /doc/begin.html: -------------------------------------------------------------------------------- 1 | 2 | 5 | 6 | 7 | 8 | 9 | CoLoR: a Coq Library on Rewriting and termination 10 | 11 | 12 | 13 | 14 |

CoLoR : a 15 | Coq 16 | Library 17 | on 18 | Rewriting and termination


19 | 20 |

Index

21 | -------------------------------------------------------------------------------- /Util/Relation/NotSN.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-08-08 6 | 7 | properties of ~SN terms (using classical logic) 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import Relations. 13 | From CoLoR Require Import ClassicUtil SN LogicUtil. 14 | 15 | Section S. 16 | 17 | Variables (A : Type) (R : relation A) (x : A) (h : ~SN R x). 18 | 19 | Lemma notSN_succ : exists y, R x y /\ ~SN R y. 20 | 21 | Proof. 22 | cut (~(forall y, R x y -> SN R y)). 23 | intro. ded (not_forall_imply_exists_not H). destruct H0. 24 | ded (imply_to_and _ _ H0). exists x0. auto. 25 | eapply contraposee_inv. apply SN_intro. exact h. 26 | Qed. 27 | 28 | End S. 29 | -------------------------------------------------------------------------------- /Util/List/ListSort.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2008-09-17 6 | 7 | sorting functions on lists 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import List. 13 | 14 | Section S. 15 | 16 | Variables (A : Type) (cmp : A -> A -> comparison). 17 | 18 | Fixpoint insert x (l : list A) := 19 | match l with 20 | | nil => cons x nil 21 | | cons y m => 22 | match cmp x y with 23 | | Gt => cons y (insert x m) 24 | | _ => cons x l 25 | end 26 | end. 27 | 28 | Fixpoint sort (l : list A) : list A := 29 | match l with 30 | | nil => nil 31 | | cons x m => insert x (sort m) 32 | end. 33 | 34 | End S. 35 | -------------------------------------------------------------------------------- /Util/Pair/PairUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-06-17 6 | 7 | general results on pairs 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil. 13 | 14 | Section S. 15 | 16 | Variables (A B : Type) 17 | (eqdecA : forall x y : A, {x=y}+{~x=y}) 18 | (eqdecB : forall x y : B, {x=y}+{~x=y}). 19 | 20 | Lemma eq_pair_dec : forall x y : A*B, {x=y}+{~x=y}. 21 | 22 | Proof. 23 | intros (x1,x2) (y1,y2). case (eqdecA x1 y1); intro. 24 | subst y1. case (eqdecB x2 y2); intro. subst y2. auto. 25 | right. unfold not. intro. injection H. intro. cong. 26 | right. unfold not. intro. injection H. intros. cong. 27 | Qed. 28 | 29 | End S. 30 | -------------------------------------------------------------------------------- /Util/Multiset/MultisetNat.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2008-10-17 6 | 7 | multisets of natural numbers 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import Arith. 13 | From CoLoR Require Import MultisetTheory MultisetOrder MultisetList RelExtras. 14 | 15 | Module Nat. 16 | Definition A := nat. 17 | End Nat. 18 | 19 | Module NatSet <: Eqset := Eqset_def Nat. 20 | 21 | Module NatSet_dec. 22 | 23 | Module Export Eq := NatSet. 24 | 25 | Definition eqA_dec := eq_nat_dec. 26 | 27 | End NatSet_dec. 28 | 29 | Module MSetCore := MultisetList.MultisetList NatSet_dec. 30 | 31 | Module Export MSetTheory := MultisetTheory.Multiset MSetCore. 32 | 33 | Module Export MSetOrd := MultisetOrder.MultisetOrder MSetCore. 34 | -------------------------------------------------------------------------------- /ProofChecker/EmptyChecker.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2009-03-24 6 | 7 | A termination solver that knows how to deal with empty problems. 8 | *) 9 | 10 | From CoLoR Require Import ListUtil Problem ATrs SN LogicUtil. 11 | 12 | Set Implicit Arguments. 13 | 14 | Section EmptySolver. 15 | 16 | Variable Sig : Signature. 17 | 18 | Definition is_problem_empty (Pb : Problem Sig) := 19 | match Pb with 20 | | FullTerm R => is_empty R 21 | | RelTerm T R => is_empty R 22 | | RelTopTerm T R => is_empty R 23 | end. 24 | 25 | Lemma empty_solver Pb : is_problem_empty Pb = true -> Prob_WF Pb. 26 | Proof. 27 | unfold Prob_WF; destruct Pb; simpl; intros; destruct R; 28 | try solve [termination_trivial | discr]. 29 | Qed. 30 | 31 | End EmptySolver. 32 | 33 | #[global] Hint Resolve empty_solver : rainbow. 34 | -------------------------------------------------------------------------------- /Coccinelle/README: -------------------------------------------------------------------------------- 1 | This is the Coccinelle_for_CoLoR library version 1.0.0. 2 | Check http://color.inria.fr/ for more recent versions. 3 | 4 | This is a modified version of Coccinelle-v8.2-bool obtained from 5 | http://a3pat.ensiie.fr/pub/coccinelle-8.2.tar.gz on 26 October 2009. 6 | 7 | The files have been slightly modified by: 8 | - removing the directory ordinals 9 | - removing the files in the directory examples (but not the directory 10 | examples and its subdirectory cime_trace) 11 | - removing all "Add LoadPath" instructions 12 | - defining every thing in Type instead of Set 13 | - replaced Add Relation/Morphism by Instance 14 | - removed Emacs local variables 15 | - removed commented code 16 | - adapting proof scripts to new Coq versions 17 | 18 | This also includes Sorin Stratulat's patch on the files 19 | term_orderings/rpo.v and examples/cime_trace/rpo_extension.v 20 | extending RPO precedence from an ordering to a quasi-ordering. 21 | -------------------------------------------------------------------------------- /ProofChecker/Problem.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2009-03-24 6 | 7 | A specification of a termination problem. 8 | *) 9 | 10 | From Stdlib Require Import List Relations. 11 | From CoLoR Require Import ATrs SN. 12 | 13 | Set Implicit Arguments. 14 | 15 | Section TerminationProblem. 16 | 17 | Variable Sig : Signature. 18 | 19 | Notation term := (term Sig). 20 | Notation rule := (rule Sig). 21 | Notation trs := (list rule). 22 | 23 | Inductive Problem : Type := 24 | | FullTerm (R : trs) 25 | | RelTerm (T R : trs) 26 | | RelTopTerm (T R : trs). 27 | 28 | Definition prob_red (P : Problem) : relation term := 29 | match P with 30 | | FullTerm R => red R 31 | | RelTerm T R => red_mod T R 32 | | RelTopTerm T R => hd_red_mod T R 33 | end. 34 | 35 | Definition Prob_WF (P : Problem) := WF (prob_red P). 36 | 37 | End TerminationProblem. 38 | -------------------------------------------------------------------------------- /Term/Varyadic/VSubstitution.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-12-05 6 | 7 | substitutions 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import Relations. 13 | From CoLoR Require Import LogicUtil VTerm. 14 | 15 | Section S. 16 | 17 | Variable Sig : Signature. 18 | 19 | Notation term := (term Sig). Notation terms := (list term). 20 | 21 | Definition substitution := variable -> term. 22 | 23 | Fixpoint sub (s : substitution) (t : term) : term := 24 | match t with 25 | | Var x => s x 26 | | Fun f ts => Fun f (map (sub s) ts) 27 | end. 28 | 29 | Lemma sub_fun : forall s f ts, sub s (Fun f ts) = Fun f (map (sub s) ts). 30 | 31 | Proof. refl. Qed. 32 | 33 | Section properties. 34 | 35 | Variable succ : relation term. 36 | 37 | Definition substitution_closed := 38 | forall t1 t2 s, succ t1 t2 -> succ (sub s t1) (sub s t2). 39 | 40 | End properties. 41 | 42 | End S. 43 | -------------------------------------------------------------------------------- /Util/Integer/BigZUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2009-03-18 6 | 7 | extension of BigZ 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil BigNUtil NatUtil. 13 | From Bignums Require Export BigZ. 14 | From Stdlib Require Import ZArith. 15 | 16 | Open Scope bigZ_scope. 17 | 18 | Lemma compare_antisym : forall x y, CompOpp (x?=y) = (y?=x). 19 | 20 | Proof. intros x y. rewrite !BigZ.spec_compare. apply Zcompare_antisym. Qed. 21 | 22 | From CoLoR Require Import OrdUtil. 23 | 24 | Lemma compare_antisym_eq : forall x y c, (x?=y = CompOpp c) <-> (y?=x = c). 25 | 26 | Proof. 27 | intros. rewrite <- (compare_antisym y x). split; intro. 28 | rewrite CompOpp_eq in H. hyp. rewrite H. refl. 29 | Qed. 30 | 31 | Lemma le_gt_dec : forall n m, {n <= m} + {n > m}. 32 | 33 | Proof. 34 | intros. unfold BigZ.le, BigZ.lt. 35 | destruct (Z_le_gt_dec (BigZ.to_Z n) (BigZ.to_Z m)). 36 | left. hyp. right. unfold Z.lt. rewrite <- Zcompare_antisym, g. refl. 37 | Defined. 38 | -------------------------------------------------------------------------------- /Term/String/SContext.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-12-05 6 | 7 | one-hole contexts 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil. 13 | From CoLoR Require Export VSignature. 14 | From Stdlib Require Export List. 15 | 16 | Section S. 17 | 18 | Variable Sig : Signature. 19 | 20 | Notation string := (list Sig). 21 | 22 | (***********************************************************************) 23 | (** contexts and replacement of the hole *) 24 | 25 | Record context : Type := mkContext { lft : string; rgt : string }. 26 | 27 | Definition fill c s := lft c ++ s ++ rgt c. 28 | 29 | (***********************************************************************) 30 | (** context composition *) 31 | 32 | Definition comp c d := mkContext (lft c ++ lft d) (rgt d ++ rgt c). 33 | 34 | Lemma fill_fill : forall c d u, fill c (fill d u) = fill (comp c d) u. 35 | 36 | Proof. 37 | intros. destruct c. destruct d. 38 | unfold fill. simpl. rewrite !app_assoc. refl. 39 | Qed. 40 | 41 | End S. 42 | 43 | Arguments mkContext [Sig] _ _. 44 | -------------------------------------------------------------------------------- /ProofChecker/Proof.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2009-03-24 6 | 7 | A specification of termination proofs. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import ZArith List. 13 | From CoLoR Require Import ASignature. 14 | 15 | Section Proof. 16 | 17 | Variable Sig : Signature. 18 | 19 | (** [trsInt si] is a representation for an interpretation assigning type 20 | [si] to a function symbol. It is represented by a list of pairs: 21 | [(f, fi)], where [f] is a function symbol and [fi] is its 22 | interpretation. *) 23 | Definition trsInt symInt := list (Sig * symInt). 24 | 25 | (** [monomial] is a monomial in representation of a polynomial. 26 | A monomial [C * x_0^i_0 * ... * x_n^i_n] is represented as: 27 | [(C, i_0::...::i_n)]. *) 28 | Definition monomial := (Z * list nat)%type. 29 | 30 | (** [polynomial] is a list of monomials so [m_0 + ... + m_n] becomes 31 | [m_0::...::m_n]. *) 32 | Definition polynomial := list monomial. 33 | 34 | Inductive TerminationProof := 35 | | TP_PolyInt (PI: trsInt polynomial) (Prf : TerminationProof) 36 | | TP_ProblemEmpty. 37 | 38 | End Proof. 39 | -------------------------------------------------------------------------------- /Coccinelle/Makefile: -------------------------------------------------------------------------------- 1 | # CoLoR, a Coq library on rewriting and termination. 2 | # See the COPYRIGHTS and LICENSE files. 3 | # 4 | # - Frederic Blanqui, 2015-02-04 5 | 6 | LIBNAME := Coccinelle_for_CoLoR 7 | 8 | .SUFFIXES: 9 | 10 | .PHONY: default config clean clean-all clean-doc tags install doc 11 | 12 | MAKECOQ := $(MAKE) -r -f Makefile.coq 13 | 14 | VFILES := $(shell find . -name \*.v | grep -v .\#) 15 | 16 | default: Makefile.coq 17 | +$(MAKECOQ) 18 | 19 | config Makefile.coq: 20 | coq_makefile -R . $(LIBNAME) $(VFILES) > Makefile.coq 21 | $(MAKECOQ) depend 22 | 23 | clean: 24 | rm -f `find . -name \*~` 25 | $(MAKECOQ) clean 26 | 27 | clean-all: clean 28 | rm -f Makefile.coq 29 | 30 | clean-doc: 31 | rm -f doc/$(LIBNAME).*.html doc/index.html doc/main.html 32 | 33 | tags: 34 | coqtags `find . -name \*.v` 35 | 36 | INSTALL_DIR ?= `coqtop -where`/user-contrib 37 | 38 | install: 39 | +for i in `find . -name \*.vo`; do \ 40 | install -d $(INSTALL_DIR)/$(LIBNAME)/$$i; \ 41 | install -m 0644 $$i $(INSTALL_DIR)/$(LIBNAME)/$$i; \ 42 | done 43 | 44 | doc: 45 | coqdoc --html -g -d doc -R . $(LIBNAME) `find . -name \*.v -print` 46 | ./createIndex 47 | 48 | %.vo: %.v 49 | $(MAKECOQ) $@ 50 | 51 | %: 52 | $(MAKECOQ) $@ 53 | -------------------------------------------------------------------------------- /stat_color: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2007-04-17 7 | 8 | prog=`basename $0` 9 | 10 | usage () { 11 | echo "usage: $prog [-h]" 12 | } 13 | 14 | help () { 15 | cat < False 23 | | _ => True 24 | end. 25 | 26 | Lemma notvar_elim : forall t, 27 | notvar t -> exists f : Sig, exists ts, t = Fun f ts. 28 | 29 | Proof. 30 | intro t. case t; simpl; intros. contr. exists f. exists t0. refl. 31 | Qed. 32 | 33 | Lemma notvar_var : forall v, ~ notvar (Var v). 34 | 35 | Proof. 36 | auto. 37 | Qed. 38 | 39 | Lemma notvar_sub : forall s t, notvar t -> notvar (sub s t). 40 | 41 | Proof. 42 | intros s t. case t; simpl; intros. contr. exact I. 43 | Qed. 44 | 45 | Lemma notvar_fill : forall c t, notvar t -> notvar (fill c t). 46 | 47 | Proof. 48 | intro c. case c; simpl; intros. exact H. exact I. 49 | Qed. 50 | 51 | Lemma notvar_fillsub : forall c s t, notvar t -> notvar (fill c (sub s t)). 52 | 53 | Proof. 54 | intros. apply notvar_fill. apply notvar_sub. hyp. 55 | Qed. 56 | 57 | End S. 58 | -------------------------------------------------------------------------------- /stat_coq: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2007-04-19 7 | 8 | prog=`basename $0` 9 | 10 | usage () { echo "usage: $prog [-h] [dir]"; } 11 | 12 | help () { 13 | cat < 32 | match PolyChecker.polySolver PI Pb with 33 | | None => Error 34 | | Some Pb' => 35 | match check_proof Pb' Prf' with 36 | | Error => Error 37 | | TerminationEstablished _ => TerminationEstablished _ 38 | end 39 | end 40 | 41 | | TP_ProblemEmpty _ => 42 | match EmptyChecker.is_problem_empty Pb with 43 | | true => TerminationEstablished _ 44 | | _ => Error 45 | end 46 | end. 47 | 48 | Next Obligation. 49 | Proof. 50 | auto with rainbow. 51 | Qed. 52 | 53 | End ProofChecker. 54 | -------------------------------------------------------------------------------- /create_dist: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2005-02-03 7 | 8 | prog=`basename $0` 9 | 10 | usage () { echo "usage: $prog [-h] [version_name]"; } 11 | 12 | help () { 13 | cat < $dir/README.md 59 | create_tar_gz 60 | md5sum $cwd/$dir.tar.gz 61 | 62 | # CoLoR with Coccinelle 63 | 64 | dir=color_with_coccinelle.$version 65 | rm -rf $dir 66 | cp -r color $dir 67 | create_tar_gz 68 | 69 | rm -rf color 70 | -------------------------------------------------------------------------------- /Coccinelle/INSTALL: -------------------------------------------------------------------------------- 1 | Compilation and installation procedure 2 | -------------------------------------- 3 | 4 | The compilation procedure that is provided hereafter requires a 5 | UNIX-like operating system with the following commands (it also works 6 | with Cygwin on Windows): 7 | 8 | - Coq version 8.4pl5 http://coq.inria.fr/ 9 | - GNU make version 3.81 http://www.gnu.org/software/make/ 10 | - GNU find version 4.4.2 http://www.gnu.org/software/findutils/findutils.html 11 | 12 | For compiling the library, do 'make [-j N]' 13 | (it takes 6'37" on my computer with N=10). 14 | 15 | If you use the development version under SVN, before doing 'make', you 16 | need to do 'make config' to recompile dependencies when new files are 17 | added in the archive. 18 | 19 | For installing the library in the default Coq user-contrib directory, do: 20 | 'make install'. 21 | 22 | For installing the library in $dir, do: 'DSTROOT=$dir make install'. 23 | 24 | For using the library without installing it, you need to type the following 25 | command in Coq or add it in your ~/.coqrc file: 26 | 27 | Add Rec LoadPath "$dir" as Coccinelle_for_CoLoR. 28 | 29 | where $dir is the installation directory. 30 | 31 | For compiling the html documentation, do 'make doc'. The starting 32 | file is doc/main.html. 33 | 34 | For compiling the TAGS file, do 'make tags'. This requires 'coqtags' 35 | from ProofGeneral (http://proofgeneral.inf.ed.ac.uk/). 36 | 37 | If you encounter any problem, send a mail to Frédéric Blanqui 38 | (firstname.name@inria.fr). 39 | -------------------------------------------------------------------------------- /Util/Relation/RedLength.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-08-08 6 | 7 | maximal reduction length of a term 8 | for a finitely branching well-founded relation 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | From CoLoR Require Import SN ListMax RelUtil ListUtil LogicUtil. 14 | From Stdlib Require Import Arith. 15 | 16 | Section S. 17 | 18 | Variables (A : Type) (R : relation A) (FB : finitely_branching R) (WF : WF R). 19 | 20 | Definition len_aux x (len : forall y, R x y -> nat) := 21 | let f i (hi : i < length (sons FB x)) := 22 | len (ith hi) (in_sons_R FB (ith_In hi)) 23 | in S (lmax (pvalues f)). 24 | 25 | Lemma len_aux_ext : forall x (f g : forall y, R x y -> nat), 26 | (forall y (p : R x y), f y p = g y p) -> len_aux f = len_aux g. 27 | 28 | Proof. 29 | intros. unfold len_aux. apply (f_equal (fun l => S (lmax l))). 30 | apply pvalues_eq. intros. apply H. 31 | Qed. 32 | 33 | Definition len : A -> nat := Fix (WF_wf_transp WF) _ len_aux. 34 | 35 | Lemma len_eq : forall x, len x = S (lmax (map len (sons FB x))). 36 | 37 | Proof. 38 | intro. unfold len. rewrite Fix_eq. 2: exact len_aux_ext. fold len. 39 | unfold len_aux. apply (f_equal (fun l => S (lmax l))). 40 | rewrite <- (@pvalues_ith_eq_map A nat). apply pvalues_eq. intros. refl. 41 | Qed. 42 | 43 | Lemma R_len : forall x y, R x y -> len x > len y. 44 | 45 | Proof. 46 | intros. rewrite len_eq. unfold gt, lt. apply le_n_S. apply in_lmax. 47 | apply in_map. apply R_in_sons. exact H. 48 | Qed. 49 | 50 | End S. 51 | -------------------------------------------------------------------------------- /.travis.yml: -------------------------------------------------------------------------------- 1 | os: linux 2 | dist: bionic 3 | language: shell 4 | 5 | services: 6 | - docker 7 | 8 | env: 9 | global: 10 | - NJOBS="2" 11 | - PACKAGE_NAME="coq-color" 12 | jobs: 13 | - COQ_IMAGE="coqorg/coq:8.11" 14 | - COQ_IMAGE="coqorg/coq:8.12" 15 | - COQ_IMAGE="coqorg/coq:dev" 16 | 17 | install: | 18 | # Prepare the COQ container 19 | docker pull ${COQ_IMAGE} 20 | docker run -d -i --init --name=COQ -v ${TRAVIS_BUILD_DIR}:/home/coq/${PACKAGE_NAME} -w /home/coq/${PACKAGE_NAME} ${COQ_IMAGE} 21 | docker exec COQ /bin/bash --login -c " 22 | # This bash script is double-quoted to interpolate Travis CI env vars: 23 | echo \"Build triggered by ${TRAVIS_EVENT_TYPE}\" 24 | export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m ' 25 | set -ex # -e = exit on failure; -x = trace for debug 26 | opam update -y 27 | opam pin add ${PACKAGE_NAME} . --kind=path --no-action -y 28 | opam config list; opam repo list; opam pin list; opam list 29 | " install 30 | 31 | script: 32 | - echo -e "${ANSI_YELLOW}Building ${PACKAGE_NAME}...${ANSI_RESET}" && echo -en 'travis_fold:start:script\\r' 33 | - | 34 | docker exec COQ /bin/bash --login -c " 35 | export PS4='+ \e[33;1m(\$0 @ line \$LINENO) \$\e[0m '; set -ex 36 | sudo chown -R coq:coq /home/coq/${PACKAGE_NAME} 37 | # First install the dependencies 38 | opam install ${PACKAGE_NAME} --deps-only -y 39 | opam list 40 | # Then install the package itself in verbose mode 41 | opam install ${PACKAGE_NAME} -v -y 42 | " script 43 | - echo -en 'travis_fold:end:script\\r' 44 | 45 | after_script: 46 | - docker stop COQ # optional but recommended 47 | -------------------------------------------------------------------------------- /Util/Logic/DepChoicePrf.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-08-08 6 | 7 | proof of dependent choice in classical logic + axiom of choice 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import ClassicalChoice. 13 | From CoLoR Require Import RelUtil LogicUtil. 14 | 15 | Section S. 16 | 17 | Variables (A : Type) (a : A) (R : relation A) (h : left_total R). 18 | 19 | Definition next_elt x := proj1_sig (h x). 20 | 21 | Inductive G : nat -> A -> Prop := 22 | | G0 : G 0 a 23 | | GS : forall i x, G i x -> G (S i) (next_elt x). 24 | 25 | Lemma G_is_classic_left_total : classic_left_total G. 26 | 27 | Proof. 28 | intro. elim x. exists a. exact G0. 29 | intros. destruct H. exists (next_elt x0). exact (GS H). 30 | Qed. 31 | 32 | Lemma G_is_functional : functional G. 33 | 34 | Proof. 35 | unfold functional. induction x; intros; inversion H; inversion H0. 36 | subst y. subst z. refl. rewrite (IHx _ _ H2 H5). refl. 37 | Qed. 38 | 39 | Lemma choice_imply_dep_choice : exists f, IS R f. 40 | 41 | Proof. 42 | destruct (choice G G_is_classic_left_total) as [f H]. exists f. 43 | assert (forall i x, G i x -> x = f i). intros. ded (H i). 44 | exact (G_is_functional H0 H1). 45 | assert (f 0 = a). ded (H 0). inversion H1. refl. 46 | assert (forall i, f (S i) = next_elt (f i)). induction i. 47 | ded (H 1). inversion H2. inversion H5. rewrite H1. subst x. refl. 48 | ded (H (S (S i))). inversion H2. ded (H0 _ _ H5). subst x. refl. 49 | intro. rewrite H2. exact (proj2_sig (h (f i))). 50 | Qed. 51 | 52 | End S. 53 | -------------------------------------------------------------------------------- /Coccinelle/examples/cime_trace/closure_extension.v: -------------------------------------------------------------------------------- 1 | From CoLoR Require Import closure. 2 | From Stdlib Require Import Relation_Operators. 3 | 4 | Section S. 5 | 6 | Variable A : Type. 7 | Variable R: Relation_Definitions.relation A. 8 | 9 | 10 | Lemma refl_trans_clos_ind2 : 11 | forall x (P:A -> Prop), 12 | (P x) -> 13 | (forall y, P y -> forall z, R y z -> P z) -> 14 | forall a, refl_trans_clos R x a -> P a. 15 | Proof. 16 | intros x P H H0 a H1. 17 | induction H1. 18 | exact H. 19 | induction H1. 20 | apply H0 with (y:=x);assumption. 21 | eauto. 22 | Qed. 23 | 24 | 25 | Lemma refl_trans_clos_R : 26 | forall x y z, refl_trans_clos R x y -> R y z -> 27 | refl_trans_clos R x z. 28 | Proof. 29 | intros x y z Rs_x_y R_y_z. 30 | apply refl_trans_clos_is_trans with y. 31 | assumption. 32 | do 2 constructor;assumption. 33 | Qed. 34 | 35 | Lemma refl_trans_clos_with_R : 36 | forall x y, R x y -> refl_trans_clos R x y. 37 | Proof. 38 | intros. 39 | constructor;constructor;assumption. 40 | Qed. 41 | 42 | Lemma trans_clos_left : 43 | forall R' x y, trans_clos R x y -> trans_clos (union _ R R') x y. 44 | Proof. 45 | intros R' x y H. 46 | induction H. 47 | constructor. left; assumption. 48 | constructor 2 with y. 49 | left;assumption. 50 | assumption. 51 | Qed. 52 | 53 | 54 | Lemma trans_clos_right : 55 | forall R' x y, trans_clos R x y -> trans_clos (union _ R' R) x y. 56 | Proof. 57 | intros R' x y H. 58 | induction H. 59 | constructor. right; assumption. 60 | constructor 2 with y. 61 | right;assumption. 62 | assumption. 63 | Qed. 64 | 65 | End S. 66 | -------------------------------------------------------------------------------- /create_index: -------------------------------------------------------------------------------- 1 | #!/bin/sh 2 | 3 | # CoLoR, a Coq library on rewriting and termination. 4 | # See the COPYRIGHTS and LICENSE files. 5 | # 6 | # - Frederic Blanqui, 2006-01-16 7 | 8 | prog=`basename $0` 9 | 10 | usage () { 11 | echo "usage: $prog [-h]" 12 | } 13 | 14 | help () { 15 | cat <$dir
    " >> $doc 38 | for f in `ls $dir/*.v` 39 | do 40 | if test "$root" = "" 41 | then 42 | d=$dir 43 | else 44 | d=$root.$dir 45 | fi 46 | g=`basename $f .v` 47 | echo "
  • $g
  • " >> $doc 48 | done 49 | echo "
" >> $doc 50 | } 51 | 52 | # add an entry for each $subdirs of $root 53 | 54 | prdir () { 55 | echo "
  • $root
      " >> $doc 56 | cd $root 57 | for d in $subdirs 58 | do 59 | dir=$d; prfiles 60 | done 61 | cd .. 62 | echo "
  • " >> $doc 63 | } 64 | 65 | # begin 66 | 67 | cat doc/begin.html > $doc 68 | echo "
      " >> $doc 69 | 70 | # Util 71 | 72 | root=Util 73 | subdirs=`ls -1 $root` 74 | prdir 75 | 76 | # Term 77 | 78 | root=Term 79 | subdirs=`ls -1 $root` 80 | prdir 81 | 82 | # simple directories 83 | 84 | root='' 85 | for dir in Conversion DP Filter HORPO MannaNess MatrixInt MPO NonTermin PolyInt ProofChecker RPO SemLab 86 | do 87 | prfiles 88 | done 89 | 90 | # end 91 | 92 | echo "
    " >> $doc 93 | cat doc/end.html >> $doc 94 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | # CoLoR, a Coq library on rewriting and termination. 2 | # See the COPYRIGHTS and LICENSE files. 3 | # 4 | # - Frederic Blanqui, 2005-02-03 5 | 6 | LIBNAME := CoLoR 7 | 8 | .SUFFIXES: 9 | 10 | .PHONY: default config clean clean-dep clean-all clean-doc tags doc install-doc install-dist targz 11 | 12 | MAKECOQ := +$(MAKE) -r -f rocq.mk 13 | 14 | VFILES := $(shell find . -name \*.v | grep -v .\# | sed -e 's|^./||g') 15 | 16 | default: rocq.mk 17 | $(MAKECOQ) 18 | 19 | time: rocq.mk 20 | COQC=./time_coqc make 21 | 22 | config rocq.mk: 23 | echo -R . $(LIBNAME) $(VFILES) > _CoqProject 24 | rocq makefile -f _CoqProject -o rocq.mk 25 | 26 | clean: 27 | rm -f `find . -name \*~` 28 | -$(MAKECOQ) clean 29 | rm -rf `find . -name .coq-native -o -name .\*.aux -o -name \*.cache` 30 | 31 | clean-dep: 32 | rm -f `find . -name \*.v.d` 33 | 34 | clean-all: clean clean-doc clean-dep 35 | rm -f _CoqProject rocq.mk rocq.mk.conf stat_time.log `find . -name \*.time` 36 | 37 | clean-doc: 38 | rm -f doc/$(LIBNAME).*.html doc/index.html doc/main.html doc/coqdoc.css 39 | 40 | tags: 41 | coqtags $(VFILES) 42 | 43 | doc: 44 | coqdoc --html -g -d doc -R . $(LIBNAME) `find . -path ./Coccinelle -prune -o -name \*.v -print` 45 | ./create_index 46 | 47 | WEB := scm.gforge.inria.fr:/home/groups/color/htdocs 48 | 49 | install-doc: 50 | scp -r doc/coqdoc.css doc/*.html $(WEB)/doc/ 51 | 52 | install-dist: 53 | scp CHANGES $(WEB)/CHANGES.CoLoR 54 | 55 | targz: 56 | rm -rf /tmp/color 57 | cp -r . /tmp/color 58 | make -C /tmp/color clean-all 59 | rm -rf `find /tmp/color -name .svn` 60 | (cd /tmp; tar zcf color.tar.gz color) 61 | rm -rf /tmp/color 62 | mv /tmp/color.tar.gz . 63 | 64 | %.vo: %.v 65 | $(MAKECOQ) $@ 66 | 67 | %: 68 | $(MAKECOQ) $@ 69 | -------------------------------------------------------------------------------- /COPYRIGHTS: -------------------------------------------------------------------------------- 1 | CoLoR, a Coq library on rewriting and termination 2 | ------------------------------------------------- 3 | 4 | The copyrights on the parts due to Frederic Blanqui, Sebastien 5 | Hinderer, Pierre-Yves Strub and Sidi Ould Biha belong to the "Institut 6 | National de Recherche en Informatique et Automatique" 7 | (http://www.inria.fr/). 8 | 9 | The copyrights on the parts due to Adam Koprowski belong to the "Vrije 10 | Universiteit Amsterdam" (http://www.vu.nl/) for HORPO and to the 11 | "Technische Universiteit Eindhoven" (http://w3.tue.nl/) for the other 12 | things. 13 | 14 | The copyrights on the parts due to Hans Zantema belong to the 15 | "Technische Universiteit Eindhoven" (http://w3.tue.nl/). 16 | 17 | The copyrights on the parts due to Solange Coupet-Grimal and William 18 | Delobel belong to the "Universite de Provence, Aix-Marseille 1" 19 | (http://www.up.univ-mrs.fr/), the "Universite de la Mediterrannee, 20 | Aix-Marseille 2" (http://www.mediterranee.univ-aix.fr/) and the 21 | "Centre National de la Recherche Scientifique" (http://www.cnrs.fr/). 22 | 23 | The copyrights on the parts due to Stephane Le Roux belong to the 24 | "Ecole Normale Superieure de Lyon" (http://www.ens-lyon.fr/). 25 | 26 | The copyrights on the parts due to Leo Ducas belong to the "Ecole 27 | Normale Superieure de Paris" (http://www.ens.fr/). 28 | 29 | The copyrights on the parts due to Johannes Waldmann belong to the 30 | "Hochschule fur Technik, Wirtschaft und Kultur (FH) Leipzig" 31 | (http://www.htwk-leipzig.de/). 32 | 33 | The copyrights on the parts due to Qian Wang and Lianyi Zhang belong 34 | to "Tsinghua University" (http://www.tsinghua.edu.cn/). 35 | 36 | The copyrights on the parts due to Sorin Stratulat belong to 37 | "University Paul Verlaine" (http://www.univ-metz.fr/). 38 | -------------------------------------------------------------------------------- /Util/Function/NaryFunction.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Sebastien Hinderer, 2004-04-02 6 | 7 | basic definitions on functions taking vectors as arguments 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import RelUtil VecUtil. 13 | 14 | Section S. 15 | 16 | Variables A B : Type. 17 | 18 | Definition naryFunction n := vector A n -> B. 19 | 20 | Variables (n : nat) (f : naryFunction n). 21 | 22 | Definition Vmonotone_i (Ra : relation A) (Rb : relation B) 23 | i j (H : i + S j = n) := 24 | forall vi vj, monotone Ra Rb (fun z => f (Vcast (Vapp vi (Vcons z vj)) H)). 25 | 26 | Arguments Vmonotone_i Ra Rb [i j] _. 27 | 28 | Lemma Vmonotone_i_transp Ra Rb i j (H : i + S j = n) : 29 | Vmonotone_i Ra Rb H -> Vmonotone_i (transp Ra) (transp Rb) H. 30 | 31 | Proof. intros H' vi vj. apply monotone_transp. apply H'. Qed. 32 | 33 | Definition Vmonotone Ra Rb := 34 | forall i j (H : i + S j = n), Vmonotone_i Ra Rb H. 35 | 36 | Lemma Vmonotone_transp Ra Rb : 37 | Vmonotone Ra Rb -> Vmonotone (transp Ra) (transp Rb). 38 | 39 | Proof. intros H i j H'. apply Vmonotone_i_transp. apply H. Qed. 40 | 41 | End S. 42 | 43 | Definition naryFunction1 A := naryFunction A A. 44 | Definition Vmonotone1_i A n f R := @Vmonotone_i A A n f R R. 45 | Definition Vmonotone1 A n f R := @Vmonotone A A n f R R. 46 | 47 | Arguments Vmonotone_i [A B n] _ _ _ [i j] _. 48 | 49 | Section preserv. 50 | 51 | Variables (A : Type) (P : A->Prop) (n : nat) (f : naryFunction A A n). 52 | 53 | Definition preserv := forall v, Vforall P v -> P (f v). 54 | 55 | Definition restrict (H : preserv) : naryFunction (sig P) (sig P) n := 56 | fun v => exist (H _ (Vforall_of_sig v)). 57 | 58 | End preserv. 59 | -------------------------------------------------------------------------------- /NonTermin/AVarCond.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2009-11-02 6 | 7 | violation of variable condition 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil ATrs AVariables BoolUtil EqUtil ListUtil RelUtil 13 | NatUtil APosition. 14 | 15 | Section S. 16 | 17 | Variable Sig : Signature. 18 | 19 | Notation rules := (rules Sig). 20 | 21 | Lemma var_cond_mod : forall E R : rules, 22 | ~rules_preserve_vars R -> EIS (red_mod E R). 23 | 24 | Proof. 25 | intros E R. rewrite <- brules_preserve_vars_ok, <- false_not_true. 26 | unfold brules_preserve_vars. 27 | rewrite (forallb_neg (@brule_preserve_vars_ok Sig)). 28 | intros [[l r] [h1 h2]]. simpl in *. rewrite not_incl in h2. 29 | 2: apply eq_nat_dec. 30 | destruct h2. destruct H. destruct (in_vars_subterm H). rename x0 into p. 31 | destruct (subterm_pos_elim H1). rename x0 into c. destruct a. 32 | set (s := single x l). set (f := iter l (fill (subc s c))). 33 | exists f. unfold IS. induction i; simpl in *. exists (f 0). split. 34 | apply rt_refl. exists l. exists r. exists Hole. simpl. exists s. 35 | repeat split. hyp. unfold s. rewrite sub_single_not_var. refl. hyp. 36 | rewrite H3, sub_fill. unfold s, single. simpl. 37 | rewrite (beq_refl beq_nat_ok). refl. unfold f. apply red_mod_fill. hyp. 38 | Qed. 39 | 40 | Lemma var_cond : forall R : rules, ~rules_preserve_vars R -> EIS (red R). 41 | 42 | Proof. intros. rewrite <- red_mod_empty. apply var_cond_mod. hyp. Qed. 43 | 44 | End S. 45 | 46 | Ltac var_cond Sig := 47 | (apply var_cond_mod || apply var_cond); 48 | rewrite <- (ko (@brules_preserve_vars_ok Sig)); 49 | (check_eq || fail 10 "variable condition satisfied"). 50 | -------------------------------------------------------------------------------- /Util/Relation/RelMidex.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Stephane Le Roux, 2007-02-20 6 | 7 | excluded middle and decidability for relations. 8 | *) 9 | 10 | From Stdlib Require Import Relations. 11 | 12 | From CoLoR Require Import LogicUtil. 13 | 14 | Set Implicit Arguments. 15 | 16 | Section S. 17 | 18 | Variables (A : Type) (R : relation A). 19 | 20 | Definition rel_midex := forall x y : A, R x y \/ ~R x y. 21 | 22 | Definition rel_dec := forall x y, {R x y} + {~R x y}. 23 | 24 | Lemma rel_dec_midex : rel_dec -> rel_midex. 25 | 26 | Proof. do 3 intro. destruct (X x y); tauto. Qed. 27 | 28 | Definition fun_rel_dec (f : A->A->bool) := 29 | forall x y, if f x y then R x y else ~R x y. 30 | 31 | Lemma bool_rel_dec : {f : A->A->bool | fun_rel_dec f} -> rel_dec. 32 | 33 | Proof. intros (f,H) x y. gen (H x y). case (f x y); intros; tauto. Qed. 34 | 35 | Lemma rel_dec_bool : rel_dec -> {f : A->A->bool | fun_rel_dec f}. 36 | 37 | Proof. 38 | intro H. exists (fun x y : A => if H x y then true else false). 39 | intros x y. destruct (H x y); trivial. 40 | Qed. 41 | 42 | Lemma fun_rel_dec_true : forall f x y, fun_rel_dec f -> f x y = true -> R x y. 43 | 44 | Proof. intros. set (w := H x y). rewrite H0 in w. hyp. Qed. 45 | 46 | Lemma fun_rel_dec_false : forall f x y, 47 | fun_rel_dec f -> f x y = false -> ~R x y. 48 | 49 | Proof. intros. set (w := H x y). rewrite H0 in w. hyp. Qed. 50 | 51 | (***********************************************************************) 52 | (** Leibniz equality relation *) 53 | 54 | Definition eq_midex := forall x y : A, x=y \/ x<>y. 55 | 56 | Definition eq_dec := forall x y : A, {x=y}+{x<>y}. 57 | 58 | Lemma eq_dec_midex : eq_dec -> eq_midex. 59 | 60 | Proof. do 3 intro. destruct (X x y); tauto. Qed. 61 | 62 | End S. 63 | -------------------------------------------------------------------------------- /Coccinelle/basis/terminaison.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Lia. 2 | From Stdlib Require Import Relations Wellfounded. 3 | Set Implicit Arguments. 4 | 5 | Section star. 6 | Unset Implicit Arguments. 7 | Variable A:Type. 8 | Variable R : A -> A -> Prop. 9 | Inductive star (x:A) : A -> Prop := 10 | | star_refl : star x x 11 | | star_step : forall y, star x y -> forall z, R y z -> star x z 12 | . 13 | 14 | 15 | Lemma star_trans: forall x y z, star x y -> star y z -> star x z. 16 | Proof. 17 | intros x y z H H1;revert x H; 18 | induction H1. 19 | 20 | tauto. 21 | intros. 22 | econstructor 2 with y0;auto. 23 | Qed. 24 | 25 | Lemma star_R : forall x y, R x y -> star x y. 26 | Proof. 27 | intros x y H;constructor 2 with x;[constructor|assumption]. 28 | Qed. 29 | 30 | Lemma star_ind2 : forall x (P:A -> Prop), 31 | (P x) -> 32 | (forall y, P y -> forall z, R y z -> P z) -> 33 | forall a, star x a -> P a. 34 | Proof. 35 | intros x P H H0 a H1. 36 | induction H1. 37 | exact H. 38 | apply H0 with (y:=y);assumption. 39 | Qed. 40 | 41 | End star. 42 | 43 | Ltac prove_star := 44 | match goal with 45 | | |- star _ _ ?t ?t => exact rt_refl 46 | | H:star _ _ ?t' ?t |- star _ _ ?t' ?t => assumption 47 | | H:star _ _ ?t'' ?t |- star _ _ ?t' ?t => 48 | (apply star_trans with t'';[prove_star|exact H]) || 49 | (clear H;prove_star) 50 | | H:star _ _ ?t' ?t'' |- star _ _ ?t' ?t => 51 | (apply star_trans with t'';[exact H|prove_star]) || 52 | (clear H;prove_star) 53 | | H:?R ?t' ?t |- star _ _ ?t' ?t => apply star_R;exact H 54 | | H:?R ?t'' ?t |- star _ _ ?t' ?t => 55 | (apply star_trans with t'';[prove_star|apply star_R;exact H]) || 56 | (clear H;prove_star) 57 | | H:?R ?t' ?t'' |- star _ _ ?t' ?t => 58 | (apply star_trans with t'';[apply star_R;exact H|prove_star]) || 59 | (clear H;prove_star) 60 | | _ => solve [eauto] 61 | end. 62 | 63 | 64 | -------------------------------------------------------------------------------- /INSTALL: -------------------------------------------------------------------------------- 1 | Compilation and installation procedure using OPAM 2 | ------------------------------------------------- 3 | 4 | Requirements: 5 | - Opam (https://opam.ocaml.org/) 6 | 7 | Procedure: 8 | - step 1/2: opam repo add coq-released https://coq.inria.fr/opam/released 9 | or opam update # if repository coq-released already added 10 | - step 2/2: opam install --jobs=N coq-color 11 | 12 | Compilation time: about 6'. 13 | 14 | Remark: to use a CoLoR file in your development, you need to prefix it 15 | with CoLoR and its complete path from the root directory of CoLoR. For 16 | instance, for LogicUtil, you need to write CoLoR.Util.Logic.LogicUtil. 17 | 18 | Enjoy! 19 | 20 | Self compilation and installation 21 | --------------------------------- 22 | 23 | Requirements: 24 | - UNIX-like operating system (including Cygwin) 25 | - GNU find version 4.7.0 (http://www.gnu.org/software/findutils/findutils.html) 26 | - GNU make version 4.1 (http://www.gnu.org/software/make/) 27 | - Coq >= 8.12 (http://coq.inria.fr/) 28 | - coq-bignums >= 8.12 (https://github.com/coq/bignums) 29 | 30 | Procedure for installing CoLoR in Coq default directory for user contributions: 31 | - step 1/2: make [-j N] 32 | - step 2/2: make install 33 | 34 | Procedure for installing CoLoR in $dir: 35 | - step 1/3: make [-j N] 36 | - step 2/3: DSTROOT=$dir make install 37 | - step 3/3: add in your _CoqProject file the following option: 38 | -R $dir CoLoR 39 | or add in your ~/.coqrc file the following line: 40 | Add Rec LoadPath "$dir" as CoLoR. 41 | 42 | Procedure for compiling the html documentation: 43 | - step 1/1: make doc 44 | (the starting file is doc/main.html) 45 | 46 | Procedure for compiling the TAGS file: 47 | - step 1/1: make tags 48 | This requires 'coqtags' from ProofGeneral (http://proofgeneral.inf.ed.ac.uk/). 49 | 50 | If you encounter any problem, send a mail to Frédéric Blanqui 51 | (firstname.name@inria.fr). 52 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | CoLoR, a Coq library on rewriting and termination 2 | ------------------------------------------------- 3 | 4 | This software is a Coq library (Coq is a proof assistant developped at 5 | INRIA available at http://coq.inria.fr/) providing a formalization of 6 | concepts and proofs of theorems on rewriting theory and termination. 7 | 8 | This software is governed by the CeCILL license under French law and 9 | abiding by the rules of distribution of free software. You can use, 10 | modify and/or redistribute the software under the terms of the CeCILL 11 | license as circulated by CEA, CNRS and INRIA at the following URL 12 | "http://www.cecill.info". A copy is provided with CoLoR in the files 13 | Licence_CeCILL_V2.1-en.txt (English) and Licence_CeCILL_V2.1-fr.txt 14 | (French). 15 | 16 | As a counterpart to the access to the source code and rights to copy, 17 | modify and redistribute granted by the license, users are provided only 18 | with a limited warranty and the software's author, the holder of the 19 | economic rights, and the successive licensors have only limited 20 | liability. 21 | 22 | In this respect, the user's attention is drawn to the risks associated 23 | with loading, using, modifying and/or developing or reproducing the 24 | software by the user in light of its specific status of free software, 25 | that may mean that it is complicated to manipulate, and that also 26 | therefore means that it is reserved for developers and experienced 27 | professionals having in-depth computer knowledge. Users are therefore 28 | encouraged to load and test the software's suitability as regards their 29 | requirements in conditions enabling the security of their systems and/or 30 | data to be ensured and, more generally, to use and operate it in the 31 | same conditions as regards security. 32 | 33 | The fact that you are presently reading this means that you have had 34 | knowledge of the CeCILL license and that you accept its terms. 35 | -------------------------------------------------------------------------------- /Util/Nat/LeastNat.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Sidi Ould-Biha, 2010-04-27 6 | 7 | Definitions and proofs about the min of a non empty set of natural 8 | numbers, using classical logic and the axiom of indefinite 9 | description. *) 10 | 11 | From Stdlib Require Import Classical_Prop IndefiniteDescription Wf_nat Relations. 12 | 13 | Set Implicit Arguments. 14 | 15 | (***********************************************************************) 16 | (** unary predicate *) 17 | 18 | Section P1. 19 | 20 | Variables (P : nat -> Prop) (exP : exists n, P n). 21 | 22 | Lemma ch_min_proof : has_unique_least_element le P. 23 | 24 | Proof. 25 | apply dec_inh_nat_subset_has_unique_least_element; auto. 26 | intros. apply classic. 27 | Qed. 28 | 29 | Definition ch_min := constructive_indefinite_description _ ch_min_proof. 30 | 31 | Lemma ch_minP : P (proj1_sig ch_min). 32 | 33 | Proof. destruct (proj2_sig ch_min) as [H _]. destruct H. auto. Qed. 34 | 35 | Lemma is_min_ch : forall n, P n -> proj1_sig ch_min <= n. 36 | 37 | Proof. 38 | destruct (proj2_sig ch_min) as [H _]. destruct H. intros. apply H0. auto. 39 | Qed. 40 | 41 | End P1. 42 | 43 | Arguments ch_min [P] _. 44 | 45 | (***********************************************************************) 46 | (** binary predicate *) 47 | 48 | Section P2. 49 | 50 | Variables (P2 : relation nat) (exP2 : forall m, exists n, P2 m n). 51 | 52 | Fixpoint rec_ch_min n : nat := 53 | match n with 54 | | S n' => proj1_sig (ch_min (exP2 (S (rec_ch_min n')))) 55 | | 0 => proj1_sig (ch_min (exP2 0)) 56 | end. 57 | 58 | Lemma rec_ch_minP : forall i, P2 (S (rec_ch_min i)) (rec_ch_min (S i)). 59 | 60 | Proof. intros. simpl. apply ch_minP. Qed. 61 | 62 | Lemma is_min_rec_ch : forall i n, 63 | P2 (S (rec_ch_min i)) n -> (rec_ch_min (S i)) <= n. 64 | 65 | Proof. intros. simpl. apply is_min_ch. auto. Qed. 66 | 67 | End P2. 68 | -------------------------------------------------------------------------------- /Util/Relation/NotSN_IS.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-08-08 6 | 7 | ~SN terms give infinite sequences (using classical logic and dependent choice) 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import DepChoice NotSN SN RelUtil LogicUtil ClassicUtil. 13 | 14 | Section S. 15 | 16 | Variables (A : Type) (R : relation A). 17 | 18 | Lemma notSN_NT : forall a, ~SN R a -> NT R a. 19 | 20 | Proof. 21 | intros a h. set (P := fun x => ~SN R x). set (B := sig P). 22 | set (T := Rof R (@proj1_sig A P)). assert (forall x, exists y, T x y). 23 | intro. destruct x. unfold T. simpl. ded (notSN_succ p). decomp H. 24 | exists (@exist _ P _ H2). hyp. 25 | set (b := @exist _ P _ h). ded (@dep_choice _ b _ H). destruct H0 as [g Hg]. 26 | set (f := fun x => proj1_sig (g x)). exists f. split; unfold f; auto. 27 | rewrite (proj2 Hg). refl. 28 | intro. ded (proj1 Hg i). destruct (g i). destruct (g (S i)). hyp. 29 | Qed. 30 | 31 | Lemma notNT_SN : forall a, ~NT R a -> SN R a. 32 | 33 | Proof. intros a ha. apply NNPP. intro h. ded (notSN_NT h). contr. Qed. 34 | 35 | Lemma notWF_EIS : ~WF R -> EIS R. 36 | 37 | Proof. 38 | unfold EIS, WF. rewrite not_forall_eq. intros [a h]. 39 | destruct (notSN_NT h). exists x. intuition. 40 | Qed. 41 | 42 | Lemma SN_notNT_eq : forall x, SN R x <-> ~NT R x. 43 | 44 | Proof. 45 | intro x. split. 46 | apply SN_notNT. 47 | apply contraposee. intro h. cut (NT R x). fo. revert h. apply notSN_NT. 48 | Qed. 49 | 50 | Lemma WF_notIS_eq : WF R <-> forall f, ~IS R f. 51 | 52 | Proof. 53 | split. intros wf f hf. 54 | ded (wf (f 0)). rewrite SN_notNT_eq in H. apply H. exists f. auto. 55 | intros h x. rewrite SN_notNT_eq. intros [f [h0 hf]]. eapply h. apply hf. 56 | Qed. 57 | 58 | Lemma notWF_EIS_eq : ~WF R <-> EIS R. 59 | 60 | Proof. 61 | rewrite WF_notIS_eq, not_forall_eq. 62 | intuition; destruct H as [f hf]; exists f. apply NNPP. hyp. auto. 63 | Qed. 64 | 65 | End S. 66 | -------------------------------------------------------------------------------- /ProofChecker/IntBasedChecker.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2009-03-25 6 | 7 | A module with results for solvers using interpretations. 8 | *) 9 | 10 | From Stdlib Require Import Program. 11 | From CoLoR Require Import LogicUtil OptUtil ATrs NaryFunction 12 | AInterpretation Proof ListUtil. 13 | 14 | Set Implicit Arguments. 15 | 16 | Section IntBasedSolver. 17 | 18 | Variable Sig : Signature. 19 | 20 | Variable rawSymInt : Type. 21 | 22 | Definition rawTrsInt := trsInt Sig rawSymInt. 23 | 24 | Variable arSymInt : nat -> Set. 25 | 26 | Variable defaultInt : forall n, arSymInt n. 27 | 28 | Definition funInt (f : Sig) := arSymInt (arity f). 29 | Definition symInt := { f : Sig & funInt f }. 30 | 31 | Definition defaultIntForSymbol f := @defaultInt (@arity Sig f). 32 | 33 | Definition buildSymInt f (fi : funInt f) : symInt := existT fi. 34 | 35 | Variable checkInt : Sig -> rawSymInt -> option symInt. 36 | 37 | Fixpoint processInt (ri : rawTrsInt) : option (list symInt) := 38 | match ri with 39 | | nil => Some nil 40 | | cons i is => 41 | match checkInt (fst i) (snd i) with 42 | | None => None 43 | | Some fi => 44 | match processInt is with 45 | | None => None 46 | | Some fis => Some (fi :: fis) 47 | end 48 | end 49 | end. 50 | 51 | Definition buildInt (i : list symInt) : forall f, funInt f := 52 | fun f => lookup_dep (el := f) (@eq_symb_dec Sig) defaultIntForSymbol i. 53 | 54 | Variable P : symInt -> Prop. 55 | Variable check_P : forall (i : symInt), option (P i). 56 | Variable default_P : forall f, P (buildSymInt (defaultIntForSymbol f)). 57 | 58 | Program Fixpoint checkProp (i : list symInt) : 59 | option (forall f, P (buildSymInt (buildInt i f))) := 60 | match lforall_opt P check_P i with 61 | | None => None 62 | | Some _ => Some _ 63 | end. 64 | 65 | Next Obligation. 66 | apply (lookup_dep_prop (P := fun _ fi => P (buildSymInt fi))); intros. 67 | destruct_call lforall_opt; try discr. 68 | ded (lforall_in l H). decomp H0. destruct x. hyp. 69 | apply default_P. 70 | Qed. 71 | 72 | End IntBasedSolver. 73 | -------------------------------------------------------------------------------- /description: -------------------------------------------------------------------------------- 1 | Name: CoLoR 2 | Title: A Coq library on Rewriting and termination 3 | Author: Frédéric Blanqui 4 | Author: Sébastien Hinderer 5 | Author: Pierre-Yves Strub 6 | Author: Sidi Ould Biha 7 | Institution: INRIA 8 | Author: Solange Coupet-Grimal 9 | Author: William Delobel 10 | Institution: CNRS 11 | Institution: Université de la Méditerranée Aix-Marseille 2 12 | Institution: Université de Provence Aix-Marseille 1 13 | Author: Adam Koprowski 14 | Institution: Vrije Universiteit Amsterdam 15 | Institution: Technische Universiteit Eindhoven 16 | Author: Hans Zantema 17 | Institution: Technische Universiteit Eindhoven 18 | Author: Stephane Le Roux 19 | Institution: Ecole Normale Superieure de Lyon 20 | Author: Leo Ducas 21 | Institution: Ecole Normale Superieure de Paris 22 | Author: Johannes Waldmann 23 | Institution: Hochschule fur Technik, Wirtschaft und Kultur (FH) Leipzig 24 | Author: Qian Wang 25 | Author: Lianyi Zhang 26 | Institution: Tsinghua University, Beijing, China 27 | Author: Sorin Stratulat 28 | Institution: University Paul Verlaine, Metz, France 29 | Description: The aim of this Coq library is to provide the necessary 30 | formal basis for certifying termination proofs for rewriting systems 31 | Url: http://color.inria.fr 32 | License: CeCILL 33 | Keywords: termination, rewriting, lambda-calculus, polynom, typing, 34 | term, multiset, vector, list, rpo, lpo, mpo, horpo, dependency pair, 35 | filtering, conversion, manna-ness, interpretation, context, substitution, 36 | lexicographic, ordering, well-founded, algebra, monotony, matrix, 37 | string, unification, graph, decomposition, matching, semantic labeling, 38 | universal algebra, first-order terms, varyadic terms, simply-typed lambda-terms, 39 | ring, semi-ring, dependency graph, relation, classical logic, 40 | ordered ring, ordered semi-ring, noetherian, root labeling, 41 | flat context closure, loop, finitely branching, reducts, dependent choice, 42 | finite graph, alpha-equivalence, transitive closure, Tait, Girard, 43 | reductibility, computability predicate, inductive type, fixpoint, Tarski, 44 | de Bruijn indices, strongly connected component, topological ordering, 45 | infinite sequence, cycle, path, set, finite map, finite set, path ordering, 46 | pigeon-hole principle, Ramsey theorem 47 | -------------------------------------------------------------------------------- /Util/List/ListMax.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-01-22 6 | - Sebastien Hinderer, 2004-04-02 7 | 8 | greatest/smallest component of a list of natural numbers 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | From CoLoR Require Import ListUtil LogicUtil NatUtil. 14 | 15 | Notation nats := (list nat). 16 | 17 | (***********************************************************************) 18 | (** max *) 19 | 20 | Fixpoint lmax (l : nats) : nat := 21 | match l with 22 | | nil => 0 23 | | a :: l' => max a (lmax l') 24 | end. 25 | 26 | Lemma in_lmax : forall x l, In x l -> x <= lmax l. 27 | 28 | Proof. 29 | induction l; simpl; intro. contr. 30 | destruct H. subst a. apply Nat.le_max_l. 31 | ded (IHl H). apply le_max_intro_r. exact H0. 32 | Qed. 33 | 34 | Arguments in_lmax [x l] _. 35 | 36 | Lemma incl_lmax : forall l1 l2, incl l1 l2 -> lmax l1 <= lmax l2. 37 | 38 | Proof. 39 | intros l1 l2. induction l1 as [| a' l' Hrec]. 40 | auto with arith. 41 | intro H. gen (incl_cons_l H). clear H. intros (H1, H2). simpl. 42 | apply (Nat.max_case a' (lmax l') (fun n : nat => n <= lmax l2)). 43 | apply in_lmax. hyp. 44 | apply Hrec. hyp. 45 | Qed. 46 | 47 | Lemma lmax_app m : forall l, lmax (l ++ m) = max (lmax l) (lmax m). 48 | 49 | Proof. 50 | induction l as [| a l Hrec]. auto. simpl. rewrite Hrec, max_assoc. refl. 51 | Qed. 52 | 53 | Lemma lmax_in : forall l, length l > 0 -> exists x, In x l /\ lmax l = x. 54 | 55 | Proof. 56 | induction l; simpl; intros. contradict H; lia. destruct l. 57 | exists a. intuition auto with *. set (l' := n::l) in *. assert (length l'>0). simpl. lia. 58 | ded (IHl H0). do 2 destruct H1. case (Nat.max_dec a (lmax l')); intro. 59 | exists a. intuition. exists x. intuition auto with *. 60 | Qed. 61 | 62 | (***********************************************************************) 63 | (** min *) 64 | 65 | Fixpoint lmin (l : nats) : nat := 66 | match l with 67 | | nil => 0 68 | | a :: l' => min a (lmin l') 69 | end. 70 | 71 | Lemma in_lmin : forall x l, In x l -> lmin l <= x. 72 | 73 | Proof. 74 | induction l; simpl; intro. contr. 75 | destruct H. subst a. apply Nat.le_min_l. 76 | ded (IHl H). apply elim_min_r. exact H0. 77 | Qed. 78 | 79 | Arguments in_lmin [x l] _. 80 | -------------------------------------------------------------------------------- /Util/Relation/Cycle.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-02-06 6 | 7 | cycles 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import Relations. 13 | From CoLoR Require Import LogicUtil Path ListUtil ListNodup. 14 | 15 | Section S. 16 | 17 | Variables (A : Type) (R : relation A) (eq_dec : forall x y : A, {x=y}+{~x=y}). 18 | 19 | (***********************************************************************) 20 | (** cycles *) 21 | 22 | Definition cycle x := path R x x. 23 | 24 | Lemma path_cycle : forall x y l, path R x y l -> In x l -> 25 | exists m, exists p, l = m ++ x :: p /\ cycle x m. 26 | 27 | Proof. 28 | intros. ded (in_elim H0). do 2 destruct H1. subst l. 29 | ded (path_app_elim H). destruct H1. 30 | exists x0. exists x1. intuition. 31 | Qed. 32 | 33 | Arguments path_cycle [x y l] _ _. 34 | 35 | (***********************************************************************) 36 | (** cycles of minimum length *) 37 | 38 | Definition cycle_min x l := cycle x l /\ ~In x l /\ nodup l. 39 | 40 | Lemma cycle_min_intro : forall x l, cycle x l -> 41 | exists m, exists y, exists p, exists q, 42 | x :: l = m ++ y :: p ++ q /\ cycle_min y p. 43 | 44 | Proof. 45 | intros. unfold cycle_min. ded (nodup_intro eq_dec (x::l)). decomp H0. 46 | (* nodup (x::l) *) 47 | exists nil. exists x. exists l. exists nil. rewrite app_nil_r. 48 | simpl in H1. intuition. 49 | (* x::l = x0++x1::x2 *) 50 | rewrite H1. ded (in_elim H4). decomp H0. rewrite H5. 51 | exists x3. exists x1. exists x4. exists (x1::x2). rewrite <- app_assoc. simpl. 52 | rewrite H5 in H2. ded (nodup_app_elim H2). simpl in H0. decomp H0. 53 | intuition. 54 | (* cycle x1 x4 *) 55 | destruct x0. contr. injection H1; intros. subst a. 56 | unfold cycle in H. destruct x3; injection H5; intros. 57 | (* x3=nil *) 58 | subst x1. subst x4. rewrite H0 in H. ded (path_app_elim H). intuition. 59 | (* x3=x::x3 *) 60 | subst a. subst x0. rewrite <- app_assoc in H0. simpl in H0. rewrite H0 in H. 61 | ded (path_app_elim H). destruct H7. ded (path_app_elim H10). intuition. 62 | Qed. 63 | 64 | End S. 65 | 66 | (***********************************************************************) 67 | (** implicit arguments *) 68 | 69 | Arguments cycle_min_intro [A R] _ [x l] _. 70 | -------------------------------------------------------------------------------- /RPO/VPrecedence.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Solange Coupet-Grimal and William Delobel, 2006-01-09 6 | 7 | recursive path orderings are monotonic well-founded strict orders 8 | *) 9 | 10 | From CoLoR Require Import VSignature VTerm RelMidex Preorder. 11 | From Stdlib Require Import Setoid. 12 | From CoLoR Require RelExtras. 13 | 14 | Module Type VPrecedenceType. 15 | 16 | Parameter Sig : Signature. 17 | 18 | Notation term := (term Sig). 19 | Notation terms := (list term). 20 | 21 | Parameter leF : Sig -> Sig -> Prop. 22 | 23 | (***********************************************************************) 24 | (** precedence *) 25 | 26 | Definition ltF := ltA Sig leF. 27 | Definition eqF := eqA Sig leF. 28 | 29 | Parameter ltF_wf : well_founded ltF. 30 | Parameter leF_dec : rel_dec leF. 31 | Parameter leF_preorder : preorder Sig leF. 32 | 33 | Infix "=F=" := eqF (at level 50). 34 | Infix " nat -> Set) in exact Set | exact Type]. 39 | Defined. 40 | 41 | Variables (w : univ_of_cycles) (eq_dec : forall x y : w, {x=y}+{~x=y}). 42 | 43 | Lemma eq_word_dec : forall n, forall x y : word w n, {x=y}+{~x=y}. 44 | 45 | Proof. induction n; simpl. auto. apply eq_zn2z_dec. hyp. Defined. 46 | 47 | End word. 48 | 49 | From CoLoR Require Import EqUtil. 50 | 51 | (***********************************************************************) 52 | (** properties of ?= on BigN *) 53 | Import ZArith. 54 | 55 | Open Scope bigN_scope. 56 | 57 | Lemma compare_antisym : forall x y, CompOpp (x?=y) = (y?=x). 58 | 59 | Proof. intros x y. rewrite !spec_compare. apply Zcompare_antisym. Qed. 60 | 61 | From CoLoR Require Import OrdUtil. 62 | 63 | Lemma compare_antisym_eq : forall x y c, (x?=y = CompOpp c) <-> (y?=x = c). 64 | 65 | Proof. 66 | intros. rewrite <- (compare_antisym y x). split; intro. 67 | rewrite CompOpp_eq in H. hyp. rewrite H. refl. 68 | Qed. 69 | 70 | (***********************************************************************) 71 | (** decidability of comparison *) 72 | 73 | Lemma bigN_le_gt_dec : forall n m, {n <= m} + {n > m}. 74 | 75 | Proof. 76 | intros. unfold le, lt. destruct (Z_le_gt_dec [n] [m]). left. hyp. right. 77 | unfold Z.lt. rewrite <- Zcompare_antisym, g. refl. 78 | Defined. 79 | -------------------------------------------------------------------------------- /Util/Vector/VecMax.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-02-02 6 | - Sebastien Hinderer, 2004-04-29 7 | 8 | greatest/smallest component of a vector of natural numbers 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | From CoLoR Require Import VecUtil NatUtil LogicUtil. 14 | 15 | Notation nats := (vector nat). 16 | 17 | (***********************************************************************) 18 | (** max *) 19 | 20 | Fixpoint Vmax n (v : nats n) : nat := 21 | match v with 22 | | Vnil => O 23 | | Vcons a w => max a (Vmax w) 24 | end. 25 | 26 | Lemma Vmax_in : forall x n (v : nats n), Vin x v -> x <= Vmax v. 27 | 28 | Proof. 29 | induction v; simpl; intros. contr. destruct H. subst h. 30 | apply le_max_intro_l. apply Nat.le_refl. apply le_max_intro_r. auto. 31 | Qed. 32 | 33 | Arguments Vmax_in [x n v] _. 34 | 35 | Lemma Vmax_head : forall n (v : nats (S n)), Vhead v <= Vmax v. 36 | 37 | Proof. intros n v. rewrite (VSn_eq v). simpl. auto with arith. Qed. 38 | 39 | Lemma Vmax_tail : forall n (v : nats (S n)), Vmax (Vtail v) <= Vmax v. 40 | 41 | Proof. intros n v. rewrite (VSn_eq v). simpl. auto with arith. Qed. 42 | 43 | Lemma Vmax_app_cons : forall n1 (v1 : nats n1) p n2 (v2 : nats n2), 44 | p <= Vmax (Vapp v1 (Vcons p v2)). 45 | 46 | Proof. 47 | intros. elim v1. simpl. auto with arith. 48 | intros a n' w H. simpl. eapply Nat.le_trans. apply H. auto with arith. 49 | Qed. 50 | 51 | Lemma Vmax_forall : forall n (v : nats n) p, 52 | Vmax v <= p -> Vforall (fun n => n <= p) v. 53 | 54 | Proof. 55 | intros n v p. elim v. intro. simpl. auto. 56 | simpl. intros h n' t Hrec H. split. 57 | eapply Nat.le_trans. apply (Nat.le_max_l h (Vmax t)). hyp. 58 | apply (Hrec (Nat.le_trans (Nat.le_max_r h (Vmax t)) H)). 59 | Qed. 60 | 61 | Arguments Vmax_forall [n v p] _. 62 | 63 | Lemma Vmax_cast : forall n (v : nats n) p (e : n=p), Vmax (Vcast v e) = Vmax v. 64 | 65 | Proof. 66 | induction v; destruct p; intros; try lia. 67 | rewrite Vcast_refl. refl. rewrite Vcast_cons. simpl. rewrite IHv. refl. 68 | Qed. 69 | 70 | (***********************************************************************) 71 | (** min *) 72 | 73 | Fixpoint Vmin n (v : nats n) : nat := 74 | match v with 75 | | Vnil => O 76 | | Vcons a w => min a (Vmin w) 77 | end. 78 | -------------------------------------------------------------------------------- /Util/Relation/SCC.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Leo Ducas, 2007-08-06 6 | 7 | Strongly Connected Components (SCC) of a graph seen as a relation 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import Cycle Path ListUtil RelUtil LogicUtil. 13 | 14 | Section S. 15 | 16 | Variable A : Type. 17 | 18 | Section definition. 19 | 20 | Variable R : relation A. 21 | 22 | (** Definition of SCC seen as a relation : are x and y in the same SCC ? *) 23 | 24 | Definition SCC x y := R! x y /\ R! y x. 25 | 26 | (** Basic properties *) 27 | 28 | Lemma SCC_trans : forall x y z, SCC x y -> SCC y z -> SCC x z. 29 | 30 | Proof. 31 | intros. unfold SCC in *. destruct H; destruct H0. 32 | split; eapply t_trans; eauto; auto. 33 | Qed. 34 | 35 | Lemma SCC_sym : forall x y, SCC x y -> SCC y x. 36 | 37 | Proof. intros. unfold SCC in *. destruct H; split; hyp. Qed. 38 | 39 | Lemma cycle_in_SCC : forall x l, cycle R x l -> forall y, In y l -> SCC x y. 40 | 41 | Proof. 42 | intros. unfold SCC. unfold cycle in H. 43 | gen (in_elim H0); intros. 44 | destruct H1; destruct H1; subst. 45 | apply path_app_elim in H. 46 | destruct H. 47 | apply path_clos_trans in H; apply path_clos_trans in H1. 48 | split; auto. 49 | Qed. 50 | 51 | Lemma SCC_in_cycle : forall x y, SCC x y -> exists l, cycle R x l /\ In y l. 52 | 53 | Proof. 54 | intros. 55 | destruct H. 56 | apply clos_trans_path in H; apply clos_trans_path in H0. 57 | destruct H; destruct H0. 58 | exists (x0++y::x1). 59 | split; auto with *. 60 | unfold cycle. 61 | apply path_app; auto. 62 | Qed. 63 | 64 | Lemma cycle_in_SCC_bound : forall x l, cycle R x l -> SCC x x. 65 | 66 | Proof. 67 | intros; unfold SCC; unfold cycle in H. 68 | split; apply path_clos_trans in H; auto. 69 | Qed. 70 | 71 | End definition. 72 | 73 | Section inclusion. 74 | 75 | Variables R1 R2 : relation A. 76 | 77 | Lemma SCC_incl : R1 << R2 -> SCC R1 << SCC R2. 78 | 79 | Proof. 80 | intros. 81 | unfold inclusion; unfold SCC. 82 | intros. 83 | destruct H0. 84 | assert (R1! << R2!). 85 | apply tc_incl; hyp. 86 | split; unfold inclusion in H2. 87 | apply (H2 x y); hyp. 88 | apply (H2 y x); hyp. 89 | Qed. 90 | 91 | End inclusion. 92 | 93 | End S. 94 | -------------------------------------------------------------------------------- /Util/Vector/VecBool.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-06-08 6 | 7 | useful definitions and lemmas on boolean vectors 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import VecUtil LogicUtil. 13 | 14 | Notation bools := (vector bool). 15 | 16 | Fixpoint Vtrue n (bs : bools n) : nat := 17 | match bs with 18 | | Vnil => 0 19 | | Vcons true bs' => S (Vtrue bs') 20 | | Vcons false bs' => Vtrue bs' 21 | end. 22 | 23 | Lemma Vtrue_app : forall n1 (bs1 : bools n1) n2 (bs2 : bools n2), 24 | Vtrue (Vapp bs1 bs2) = Vtrue bs1 + Vtrue bs2. 25 | 26 | Proof. 27 | induction bs1; simpl. auto. intros. ded (IHbs1 n2 bs2). rewrite H. 28 | case h; refl. 29 | Qed. 30 | 31 | Lemma Vtrue_break : forall n1 n2 (bs : bools (n1+n2)), 32 | Vtrue (fst (Vbreak bs)) + Vtrue (snd (Vbreak bs)) = Vtrue bs. 33 | 34 | Proof. 35 | induction n1; simpl; intros. refl. VSntac bs. simpl. case (Vhead bs); simpl. 36 | f_equal. apply IHn1. apply IHn1. 37 | Qed. 38 | 39 | Arguments Vtrue_break [n1 n2] _. 40 | 41 | Lemma Vtrue_cast : forall n (bs : bools n) p (h:n=p), 42 | Vtrue (Vcast bs h) = Vtrue bs. 43 | 44 | Proof. 45 | induction bs; induction p; intros. rewrite Vcast_refl. refl. 46 | discr. discr. rewrite Vcast_cons. simpl. case h. 47 | f_equal. apply IHbs. apply IHbs. 48 | Qed. 49 | 50 | Definition Vtrue_cons_if (b : bool) n (bs : bools (S n)) := 51 | if b then S (Vtrue (Vtail bs)) else Vtrue (Vtail bs). 52 | 53 | Definition Vtrue_cons n (bs : bools (S n)) := Vtrue_cons_if (Vhead bs) bs. 54 | 55 | Lemma Vtrue_cons_eq : forall n (bs : bools (S n)), Vtrue_cons bs = Vtrue bs. 56 | 57 | Proof. intros. VSntac bs. unfold Vtrue_cons, Vtrue_cons_if. simpl. refl. Qed. 58 | 59 | Lemma Vtrue_cons_true : forall b n (bs : bools n), 60 | b = true -> S (Vtrue bs) = Vtrue (Vcons b bs). 61 | 62 | Proof. intros. subst b. refl. Qed. 63 | 64 | Lemma Vtrue_cons_false : forall b n (bs : bools n), 65 | b = false -> Vtrue bs = Vtrue (Vcons b bs). 66 | 67 | Proof. intros. subst b. refl. Qed. 68 | 69 | Lemma Vtrue_Sn_true : forall n (bs : bools (S n)), 70 | Vhead bs = true -> S (Vtrue (Vtail bs)) = Vtrue bs. 71 | 72 | Proof. intros. VSntac bs. apply Vtrue_cons_true. hyp. Qed. 73 | 74 | Lemma Vtrue_Sn_false : forall n (bs : bools (S n)), 75 | Vhead bs = false -> Vtrue (Vtail bs) = Vtrue bs. 76 | 77 | Proof. intros. VSntac bs. apply Vtrue_cons_false. hyp. Qed. 78 | 79 | Lemma Vtrue_Sn : forall n (bs : bools (S n)), 80 | Vtrue (Vcons (Vhead bs) (Vtail bs)) = Vtrue bs. 81 | 82 | Proof. intros. rewrite <- VSn_eq. refl. Qed. 83 | -------------------------------------------------------------------------------- /Coccinelle/basis/decidable_set.v: -------------------------------------------------------------------------------- 1 | (**************************************************************************) 2 | (* * *) 3 | (* _ * The Coccinelle Library / Evelyne Contejean *) 4 | (* * CNRS-LRI-Universite Paris Sud *) 5 | (* -/@|@\- * A3PAT Project *) 6 | (* -@ | @- * *) 7 | (* -\@|@/- * This file is distributed under the terms of the *) 8 | (* -v- * CeCILL-C licence *) 9 | (* * *) 10 | (**************************************************************************) 11 | 12 | From Stdlib Require Import Relation_Definitions Setoid Bool Arith. 13 | 14 | Module Type S. 15 | 16 | Parameter A : Type. 17 | Parameter eq_bool : A -> A -> bool. 18 | Parameter eq_bool_ok : forall a1 a2, match eq_bool a1 a2 with true => a1 = a2 | false => ~ a1 = a2 end. 19 | 20 | End S. 21 | 22 | Module Type ES. 23 | 24 | Parameter A : Type. 25 | Parameter eq_A : relation A. 26 | Parameter eq_bool : A -> A -> bool. 27 | Parameter eq_bool_ok : forall a1 a2, match eq_bool a1 a2 with true => eq_A a1 a2 | false => ~eq_A a1 a2 end. 28 | 29 | Parameter eq_proof : equivalence A eq_A. 30 | 31 | Add Relation A eq_A 32 | reflexivity proved by (@equiv_refl _ _ eq_proof) 33 | symmetry proved by (@equiv_sym _ _ eq_proof) 34 | transitivity proved by (@equiv_trans _ _ eq_proof) as EQA. 35 | 36 | End ES. 37 | 38 | Module Convert (DS : S) <: 39 | ES with Definition A := DS.A 40 | with Definition eq_A := @eq DS.A. 41 | 42 | Definition A := DS.A. 43 | Definition eq_A := @eq A. 44 | Definition eq_bool := DS.eq_bool. 45 | Definition eq_bool_ok := DS.eq_bool_ok. 46 | 47 | Lemma eq_proof : equivalence A eq_A. 48 | Proof. 49 | unfold eq_A; split. 50 | intro a; reflexivity. 51 | intros a1 a2 a3 H1 H2; transitivity a2; trivial. 52 | intros a1 a2 H; symmetry; trivial. 53 | Qed. 54 | 55 | Add Relation A eq_A 56 | reflexivity proved by (Relation_Definitions.equiv_refl A eq_A eq_proof) 57 | symmetry proved by (Relation_Definitions.equiv_sym A eq_A eq_proof) 58 | transitivity proved by (Relation_Definitions.equiv_trans _ _ eq_proof) as EQA. 59 | 60 | End Convert. 61 | 62 | Lemma beq_nat_ok : forall n1 n2, if Nat.eqb n1 n2 then n1 = n2 else n1 <> n2. 63 | Proof. 64 | intros n1 n2; destruct (Nat.eqb n1 n2) eqn:E. 65 | - exact (proj1 (Nat.eqb_eq n1 n2) E). 66 | - exact (proj1 (Nat.eqb_neq n1 n2) E). 67 | Defined. 68 | -------------------------------------------------------------------------------- /Term/SimpleType/TermsSig.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2006-04-27 6 | 7 | This file presents the definition of simple types 8 | for the development of theory of simpe typed lambda-calculus. 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | Module Type BaseTypes. 14 | 15 | (* Base types; denoted as a, b, c, ... *) 16 | Parameter BaseType : Type. 17 | Implicit Types a b c : BaseType. 18 | (* Equality on base types needs to be decidable *) 19 | Parameter eq_BaseType_dec : forall (a b: BaseType), {a=b}+{~a=b}. 20 | #[global] Hint Resolve eq_BaseType_dec : terms. 21 | (* To ensure that set of base types is not empty *) 22 | Parameter baseTypesNotEmpty : BaseType. 23 | 24 | End BaseTypes. 25 | 26 | (* ================================================================== 27 | Specification of simple types. They are built using base types 28 | that are given as a parameter to this module. 29 | ================================================================== *) 30 | Module SimpleTypes (BT : BaseTypes). 31 | 32 | Export BT. 33 | 34 | (* 35 | Simple types: either basic types or arrow types A->B with A and B 36 | simple types. We denote them as A, B, C, ... 37 | *) 38 | Inductive SimpleType : Type := 39 | | BasicType(T: BaseType) 40 | | ArrowType(A B : SimpleType). 41 | Notation "x --> y" := (ArrowType x y) 42 | (at level 55, right associativity) : type_scope. 43 | Notation "# x " := (BasicType x) (at level 0) : type_scope. 44 | Implicit Types A B C : SimpleType. 45 | #[global] Hint Constructors SimpleType : terms. 46 | 47 | End SimpleTypes. 48 | 49 | (* ================================================================== 50 | Specification of signature. 51 | Signature consists of a non-empty set of function symbols 52 | (FunctionSymbol) with given types (f_type). 53 | ================================================================== *) 54 | Module Type Signature. 55 | 56 | Declare Module BT : BaseTypes. 57 | 58 | Module Export ST := SimpleTypes BT. 59 | 60 | (* Function symbols, denoted as f, g, h, ... *) 61 | Parameter FunctionSymbol : Type. 62 | Implicit Types f g h : FunctionSymbol. 63 | 64 | (* Equality on function symbols needs to be decidable *) 65 | Parameter eq_FunctionSymbol_dec : forall (f g: FunctionSymbol), 66 | {f=g} + {~f=g}. 67 | #[global] Hint Resolve eq_FunctionSymbol_dec : terms. 68 | 69 | (* To ensure that set of function symbols is not empty *) 70 | Parameter functionSymbolsNotEmpty : FunctionSymbol. 71 | 72 | (* Types for function symbols *) 73 | Parameter f_type : FunctionSymbol -> SimpleType. 74 | 75 | End Signature. 76 | -------------------------------------------------------------------------------- /Util/Logic/ClassicUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2007-08-08 6 | 7 | basic classical meta-theorems 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Export Classical Setoid. 13 | From CoLoR Require Import LogicUtil. 14 | 15 | (***********************************************************************) 16 | (** Basic meta-theorems. *) 17 | 18 | Lemma contraposee : forall P Q, (~Q -> ~P) -> P -> Q. 19 | 20 | Proof. intros. apply NNPP. intro. exact (H H1 H0). Qed. 21 | 22 | Lemma not_forall_imply_exists_not : forall A (P : A -> Prop), 23 | ~(forall x, P x) -> exists x, ~P x. 24 | 25 | Proof. 26 | intros. apply NNPP. intro. apply H. intros. elim (classic (P x)). auto. 27 | intro. absurd (exists x, ~ P x). exact H0. exists x. exact H1. 28 | Qed. 29 | 30 | Arguments not_forall_imply_exists_not [A P] _. 31 | 32 | Lemma not_forall_eq : forall (A : Type) (P : A -> Prop), 33 | ~(forall x, P x) <-> exists x, ~P x. 34 | 35 | Proof. 36 | split. apply not_forall_imply_exists_not. apply exists_not_imply_not_forall. 37 | Qed. 38 | 39 | Lemma imply_eq : forall P Q : Prop, (P -> Q) <-> (~P \/ Q). 40 | 41 | Proof. 42 | split; intro. destruct (classic P). right. auto. left. hyp. 43 | intro. destruct H. absurd P; hyp. hyp. 44 | Qed. 45 | 46 | Lemma not_not_eq : forall P : Prop, ~~P <-> P. 47 | 48 | Proof. split. apply NNPP. intros h1 h2. absurd P; hyp. Qed. 49 | 50 | Lemma not_and_eq : forall P Q : Prop, ~(P /\ Q) <-> ~P \/ ~Q. 51 | 52 | Proof. 53 | split; intro. destruct (classic P). right. intro. absurd (P/\Q). hyp. auto. 54 | left. auto. intros [h1 h2]. destruct H. absurd P; hyp. absurd Q; hyp. 55 | Qed. 56 | 57 | Lemma not_or_eq : forall P Q : Prop, ~(P \/ Q) <-> ~P /\ ~Q. 58 | 59 | Proof. 60 | split. intro. split; intro; absurd (P\/Q); auto. 61 | intros [h1 h2] [h|h]. absurd P; hyp. absurd Q; hyp. 62 | Qed. 63 | 64 | Lemma not_imply_eq : forall P Q : Prop, ~(P -> Q) <-> P /\ ~Q. 65 | 66 | Proof. 67 | intros. rewrite imply_eq. rewrite not_or_eq. rewrite not_not_eq. refl. 68 | Qed. 69 | 70 | (****************************************************************************) 71 | (** Properties of Leibniz equality on [sig]. *) 72 | 73 | From CoLoR Require Import FunUtil. 74 | 75 | Section sig. 76 | 77 | Variables (A : Type) (P : A -> Prop). 78 | 79 | Lemma sig_eq : forall x y : sig P, proj1_sig x = proj1_sig y <-> x = y. 80 | 81 | Proof. 82 | intros [x_val x] [y_val y]; simpl. split; intro e. 83 | subst y_val. rewrite (proof_irrelevance _ x y). refl. 84 | gen (f_equal (@proj1_sig _ _) e). tauto. 85 | Qed. 86 | 87 | Lemma inj_proj1_sig : injective (proj1_sig (P:=P)). 88 | 89 | Proof. intros x y e. apply sig_eq. hyp. Qed. 90 | 91 | End sig. 92 | -------------------------------------------------------------------------------- /SemLab/ARootLab.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2009-07-07 6 | 7 | root labelling (Zantema & Waldmann, RTA'07) (Sternagel & Middeldorp, RTA'08) 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import ATrs LogicUtil ListUtil VecUtil BoolUtil EqUtil ASemLab SN 13 | NatUtil RelUtil AWFMInterpretation. 14 | 15 | (***********************************************************************) 16 | (** data necessary for a root labelling *) 17 | 18 | Module Type RootLab. 19 | 20 | Parameter Sig : Signature. 21 | Parameter some_symbol : Sig. 22 | 23 | Parameter Fs : list Sig. 24 | Parameter Fs_ok : forall x : Sig, List.In x Fs. 25 | 26 | End RootLab. 27 | 28 | (***********************************************************************) 29 | (** root labelling *) 30 | 31 | Module RootSemLab (Import R : RootLab) <: FinSemLab. 32 | 33 | Notation beq_symb_ok := (@beq_symb_ok Sig). 34 | Notation eq_symb_dec := (@eq_symb_dec Sig). 35 | 36 | Notation term := (term Sig). Notation terms := (vector term). 37 | Notation rule := (rule Sig). Notation rules := (rules Sig). 38 | 39 | Definition Sig := Sig. 40 | 41 | Definition I := mkInterpretation some_symbol (fun f _ => f). 42 | 43 | Definition L (f : Sig) := vector Sig (arity f). 44 | 45 | Definition beq f g (l : L f) (m : L g) := beq_vec (@beq_symb Sig) l m. 46 | 47 | Lemma beq_ok : forall f (l1 l2 : L f), beq l1 l2 = true <-> l1 = l2. 48 | 49 | Proof. 50 | intros f l1 l2. apply beq_vec_ok. apply @beq_symb_ok. 51 | Qed. 52 | 53 | Definition pi (f : Sig) (v : vector Sig (arity f)) := v. 54 | 55 | Definition beqI (t u : term) := 56 | match t, u with 57 | | Var x, Var y => beq_nat x y 58 | | Fun f _, Fun g _ => beq_symb f g 59 | | _, _ => false 60 | end. 61 | 62 | Lemma beqI_ok : rel_of_bool beqI << IR I (@eq I). 63 | 64 | Proof. 65 | intros t u. unfold rel_of_bool. 66 | destruct t; destruct u; simpl; intros; try discr. 67 | rewrite beq_nat_ok in H. subst. intro. refl. 68 | rewrite beq_symb_ok in H. subst. intro. refl. 69 | Qed. 70 | 71 | Definition Fs := Fs. 72 | Definition Fs_ok := Fs_ok. 73 | 74 | Definition Is := Fs. 75 | Definition Is_ok := Fs_ok. 76 | 77 | Definition Ls (f : Sig) := enum_tuple I Is (arity f). 78 | 79 | Lemma Ls_ok : forall f (x : L f), List.In x (Ls f). 80 | 81 | Proof. 82 | intros f x. unfold Ls. apply (enum_tuple_complete I Is_ok). 83 | Qed. 84 | 85 | End RootSemLab. 86 | 87 | Module RootLabProps (RL : RootLab). 88 | 89 | Module FSL := RootSemLab RL. 90 | 91 | Module Props := FinSemLabProps FSL. 92 | 93 | Module LabSig := Props.LabSig. 94 | 95 | Ltac rootlab := Props.semlab. 96 | 97 | End RootLabProps. 98 | -------------------------------------------------------------------------------- /Coccinelle/examples/cime_trace/rewriting.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Lia Relations Wellfounded. 2 | From Stdlib Require Inclusion. 3 | From CoLoR Require Import term. 4 | From CoLoR Require equational_theory. (* mainly for one_step...*) 5 | 6 | Notation "'SN'" := Acc. 7 | Notation "'termine'" := well_founded. 8 | 9 | Section Inclusion. 10 | Variable A : Type. 11 | Variable R1 R2 : A->A->Prop. 12 | Hypothesis Incl : inclusion _ R1 R2. 13 | 14 | Theorem SN_incl : forall x : A , SN R2 x -> SN R1 x. 15 | unfold transp. 16 | intros. 17 | eapply Inclusion.Acc_incl ; eauto. 18 | Qed. 19 | 20 | Theorem termine_incl : termine R2 -> termine R1. 21 | unfold transp. 22 | intros. 23 | eapply Inclusion.wf_incl ; eauto. 24 | Qed. 25 | 26 | End Inclusion. 27 | 28 | Section Terminaison. 29 | Variable T:Type. 30 | Variable R: T -> T -> Prop. 31 | 32 | Declare Scope rewriting. 33 | 34 | Notation "T -R> U" := (R U T) (at level 80) : rewriting. 35 | 36 | Open Scope rewriting. 37 | 38 | Definition Retoile := clos_refl_trans _ R. 39 | 40 | Theorem no_infinite_sequence2 : forall x : T , SN R x -> 41 | ~(exists f : nat -> T , f 0 = x /\ forall i : nat , f i -R> f (i+1)) . 42 | intros. 43 | induction H. 44 | intuition. 45 | elim H1. 46 | intros. 47 | apply (H0 (x0 1)). 48 | intuition. 49 | subst. 50 | apply (H4 0). 51 | exists ( fun i => (x0 (i+1))). 52 | intuition. 53 | Qed. 54 | 55 | Variable f : T -> nat. 56 | 57 | Hypothesis H_compat : forall x y: T, x -R> y -> f y < f x. 58 | 59 | Theorem termine_gt_compat : termine R. 60 | Proof. 61 | red . 62 | cut (forall n a, f a < n -> Acc R a). 63 | intros H a. 64 | apply (H (S (f a))); auto with arith. 65 | induction n. 66 | intros; absurd (f a < 0); auto with arith. 67 | intros a ltSma. 68 | apply Acc_intro. 69 | intros b ltfafb. 70 | apply IHn. 71 | assert (f b < f a). 72 | apply H_compat. 73 | trivial. 74 | lia. 75 | Qed. 76 | 77 | Inductive R2 (y x: T) : Prop := 78 | R2_intro: forall (z:T), x -R> z -> z -R> y -> R2 y x. 79 | 80 | Lemma termineR2aux : 81 | termine R -> forall x, SN R2 x /\ (forall y, R y x -> SN R2 y). 82 | intros. 83 | apply Acc_ind 84 | with (R:=R) (P:=fun x => (SN R2 x) /\ (forall y, R y x -> (SN R2 y))). 85 | intuition. 86 | apply Acc_intro. 87 | intros. 88 | elim H2. 89 | intros. 90 | generalize (H1 z H3). 91 | intuition. 92 | generalize (H1 y H2). 93 | intuition. 94 | apply H. 95 | Qed. 96 | 97 | Theorem termineR2 : termine R -> termine R2. 98 | intros rtermine. 99 | red. 100 | intros a. 101 | generalize (termineR2aux rtermine). 102 | intros hyp. 103 | decompose [and] (hyp a). 104 | auto. 105 | Qed. 106 | 107 | End Terminaison. 108 | -------------------------------------------------------------------------------- /Term/SimpleType/TermsEta.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2006-04-27 6 | 7 | The eta-reduction relation of simply typed lambda-calculus. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import RelExtras ListExtras. 13 | From Stdlib Require Import Lia. 14 | From CoLoR Require TermsBeta. 15 | 16 | Module TermsEta (Sig : TermsSig.Signature). 17 | 18 | Module Export TB := TermsBeta.TermsBeta Sig. 19 | 20 | Definition Eta_preterm Pt A := \A => prelift Pt 1 @@ %0. 21 | 22 | Definition EtaApp : forall M A B, type M = A --> B -> 23 | { M': Term | env M' = env M 24 | /\ term M' = Eta_preterm (term M) A /\ type M' = type M }. 25 | 26 | Proof. 27 | (* from M we need to construct \x. @(M', x) where M' = lift M 1 *) 28 | intros. 29 | set (RE := decl A (env M)). 30 | (* build M' *) 31 | assert (LT: RE |- prelift (term M) 1 := A --> B). 32 | replace RE with ((decl A EmptyEnv) [+] (declDummy (env M))). 33 | apply typing_ext_env_l. 34 | rewrite <- lift_term. 35 | rewrite <- H. 36 | rewrite <- (lift_type M 1). 37 | replace (declDummy (env M)) with (liftedEnv 1 (env M) 0). 38 | rewrite <- (lift_env). 39 | exact (typing (lift M 1)). 40 | unfold liftedEnv, declDummy, finalSeg; simpl. 41 | rewrite initialSeg_full; solve [trivial | lia]. 42 | rewrite env_comp_sum_comm. 43 | unfold declDummy. 44 | replace (None :: env M) 45 | with (copy (length (decl A EmptyEnv)) None ++ env M); auto. 46 | rewrite env_sum_disjoint; auto. 47 | unfold decl, declDummy. 48 | apply env_comp_cons; auto. 49 | apply env_comp_sym. 50 | apply env_comp_empty. 51 | set (L := buildT LT). 52 | (* build x *) 53 | assert (RV: RE |= 0 := A). 54 | constructor. 55 | set (RT := TVar RV). 56 | set (R := buildT RT). 57 | (* build @(M, x) *) 58 | assert (typeOk: type_left (type L) = type R). 59 | trivial. 60 | set (Mx_cond := Build_appCond L R (eq_refl (env L)) I typeOk). 61 | set (Mx := buildApp Mx_cond). 62 | assert (envOk: env Mx |= 0 := A). 63 | trivial. 64 | set (xMx_cond := Build_absCond Mx envOk). 65 | set (xMx := buildAbs xMx_cond). 66 | exists xMx; repeat split; trivial. 67 | rewrite H; trivial. 68 | Qed. 69 | 70 | Inductive EtaExpansionStep : relation Term := 71 | | EtaExp: forall M A B (Mtyp: type M = A --> B), EtaExpansionStep M (proj1_sig (EtaApp M Mtyp)). 72 | 73 | Definition EtaExpansion := Reduction EtaExpansionStep. 74 | Notation "M -e+-> N" := (EtaExpansion M N) (at level 30). 75 | 76 | Inductive EtaReductionStep : relation Term := 77 | | EtaRed: forall M A B (Mtyp: type M = A --> B), EtaReductionStep (proj1_sig (EtaApp M Mtyp)) M. 78 | 79 | Definition EtaReduction := Reduction EtaReductionStep. 80 | Notation "M -e-> N" := (EtaRed M N) (at level 30). 81 | 82 | End TermsEta. 83 | -------------------------------------------------------------------------------- /Term/String/SReverse.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2008-02-04 6 | 7 | string reversing 8 | 9 | Termination of String Rewriting Proved Automatically, H. Zantema, 10 | Journal of Automated Reasoning, 2005, volume 34, pages 105-139 11 | *) 12 | 13 | Set Implicit Arguments. 14 | 15 | From CoLoR Require Import Srs ListUtil RelUtil LogicUtil SN. 16 | 17 | Section S. 18 | 19 | Variable Sig : Signature. 20 | 21 | Definition reverse (e : rule Sig) := let (l,r) := e in mkRule (rev' l) (rev' r). 22 | 23 | Notation reverses := (List.map reverse). 24 | 25 | Lemma red_rev : forall R t u, red R t u -> red (reverses R) (rev' t) (rev' u). 26 | 27 | Proof. 28 | intros. redtac. unfold fill. rewrite !rev'_app, !app_ass. 29 | ex (rev' l) (rev' r) (mkContext (rev' (rgt c)) (rev' (lft c))). intuition. 30 | change (In (reverse (mkRule l r)) (reverses R)). apply in_map. exact H. 31 | Qed. 32 | 33 | Lemma red_rev_rtc : forall E t u, 34 | red E # t u -> red (reverses E) # (rev' t) (rev' u). 35 | 36 | Proof. 37 | intros. elim H; intros. apply rt_step. apply red_rev. exact H0. 38 | apply rt_refl. eapply rt_trans. apply H1. exact H3. 39 | Qed. 40 | 41 | Lemma red_mod_rev : forall E R t u, 42 | red_mod E R t u -> red_mod (reverses E) (reverses R) (rev' t) (rev' u). 43 | 44 | Proof. 45 | intros. redtac. exists (rev' (fill c l)). split. apply red_rev_rtc. exact H. 46 | apply red_rev. exists l. exists r. exists c. auto. 47 | Qed. 48 | 49 | Lemma WF_red_mod_rev : forall E R, 50 | WF (red_mod (reverses E) (reverses R)) -> WF (red_mod E R). 51 | 52 | Proof. 53 | unfold WF. intros E R H. cut (forall x, SN (red_mod E R) (rev' x)). 54 | intros. ded (H0 (rev' x)). rewrite rev'_rev' in H1. exact H1. 55 | intro. ded (H x). elim H0. intros. apply SN_intro. intros. 56 | ded (H2 (rev' y)). rewrite rev'_rev' in H4. apply H4. 57 | ded (red_mod_rev H3). rewrite rev'_rev' in H5. exact H5. 58 | Qed. 59 | 60 | Lemma reverse_reverse : forall a, reverse (reverse a) = a. 61 | 62 | Proof. 63 | intros [l r]. unfold reverse. rewrite !rev'_rev'. refl. 64 | Qed. 65 | 66 | Lemma reverses_reverses : forall R, reverses (reverses R) = R. 67 | 68 | Proof. 69 | intro. rewrite map_map. apply map_eq_id. intros. apply reverse_reverse. 70 | Qed. 71 | 72 | Lemma WF_red_mod_rev_eq : forall E R, 73 | WF (red_mod (reverses E) (reverses R)) <-> WF (red_mod E R). 74 | 75 | Proof. 76 | split; intro. apply WF_red_mod_rev. hyp. apply WF_red_mod_rev. 77 | rewrite !reverses_reverses. hyp. 78 | Qed. 79 | 80 | Lemma WF_red_rev_eq : forall R, WF (red (reverses R)) <-> WF (red R). 81 | 82 | Proof. 83 | intro. rewrite <- !red_mod_empty. 84 | assert (nil = reverses nil). refl. rewrite H. apply WF_red_mod_rev_eq. 85 | Qed. 86 | 87 | End S. 88 | 89 | (***********************************************************************) 90 | (** tactics for Rainbow *) 91 | 92 | Ltac rev_tac := rewrite <- WF_red_rev_eq || rewrite <- WF_red_mod_rev_eq. 93 | -------------------------------------------------------------------------------- /opam: -------------------------------------------------------------------------------- 1 | opam-version: "2.0" 2 | synopsis: "A library on rewriting theory and termination" 3 | homepage: "https://github.com/fblanqui/color/" 4 | license: "CeCILL-2.1" 5 | bug-reports: "https://github.com/fblanqui/color/issues" 6 | depends: [ 7 | "ocaml" 8 | "rocq-prover" {>= "9.0"} 9 | "rocq-bignums" {>= "9.0"} 10 | ] 11 | build: [make "-j%{jobs}%"] 12 | install: [make "-f" "rocq.mk" "install"] 13 | flags: light-uninstall 14 | remove: ["rm" "-R" "%{lib}%/coq/user-contrib/CoLoR"] 15 | maintainer: "frederic.blanqui@inria.fr" 16 | authors: [ 17 | "Frédéric Blanqui" 18 | "Adam Koprowski" 19 | "Sébastien Hinderer" 20 | "Pierre-Yves Strub" 21 | "Sidi Ould Biha" 22 | "Solange Coupet-Grimal" 23 | "William Delobel" 24 | "Hans Zantema" 25 | "Stéphane Leroux" 26 | "Léo Ducas" 27 | "Johannes Waldmann" 28 | "Qiand Wang" 29 | "Lianyi Zhang" 30 | "Sorin Stratulat" 31 | ] 32 | tags: [ 33 | "date:2025-11-11" 34 | 35 | "logpath:CoLoR" 36 | 37 | "category:Computer Science/Algorithms/Correctness proofs of algorithms" 38 | "category:Computer Science/Data Types and Data Structures" 39 | "category:Computer Science/Lambda Calculi" 40 | "category:Mathematics/Algebra" 41 | "category:Mathematics/Combinatorics and Graph Theory" 42 | "category:Mathematics/Logic/Type theory" 43 | "category:Miscellaneous/Extracted Programs/Type checking unification and normalization" 44 | 45 | "keyword:rewriting" 46 | "keyword:termination" 47 | "keyword:lambda calculus" 48 | 49 | "keyword:list" 50 | "keyword:multiset" 51 | "keyword:polynom" 52 | "keyword:vectors" 53 | "keyword:matrices" 54 | "keyword:FSet" 55 | "keyword:FMap" 56 | 57 | "keyword:term" 58 | "keyword:context" 59 | "keyword:substitution" 60 | "keyword:universal algebra" 61 | 62 | "keyword:varyadic term" 63 | "keyword:string" 64 | 65 | "keyword:alpha-equivalence" 66 | "keyword:de bruijn indices" 67 | "keyword:simple types" 68 | 69 | "keyword:matching" 70 | "keyword:unification" 71 | 72 | "keyword:relation" 73 | "keyword:ordering" 74 | "keyword:quasi-ordering" 75 | "keyword:lexicographic ordering" 76 | 77 | "keyword:ring" 78 | "keyword:semiring" 79 | 80 | "keyword:well-founded" 81 | "keyword:noetherian" 82 | "keyword:finitely branching" 83 | "keyword:dependent choice" 84 | "keyword:infinite sequences" 85 | 86 | "keyword:non-termination" 87 | "keyword:loop" 88 | 89 | "keyword:graph" 90 | "keyword:path" 91 | "keyword:transitive closure" 92 | "keyword:strongly connected component" 93 | "keyword:topological ordering" 94 | 95 | "keyword:rpo" 96 | "keyword:horpo" 97 | "keyword:dependency pair" 98 | "keyword:dependency graph" 99 | "keyword:semantic labeling" 100 | 101 | "keyword:reducibility" 102 | "keyword:Girard" 103 | 104 | "keyword:fixpoint theorem" 105 | "keyword:Tarski" 106 | 107 | "keyword:pigeon-hole principle" 108 | "keyword:Ramsey theorem" 109 | ] 110 | -------------------------------------------------------------------------------- /Util/Relation/RelSub.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Stephane Le Roux, 2007-02-20 6 | 7 | sub-relation, restriction, etc. 8 | *) 9 | 10 | From CoLoR Require Import ListShrink RelUtil ListUtil LogicUtil. 11 | 12 | Set Implicit Arguments. 13 | 14 | Section On_Rel_Gen. 15 | 16 | Variable A : Type. 17 | 18 | Section restriction. 19 | 20 | Variables (R : relation A) (l : list A). 21 | 22 | Definition restriction x y := In x l /\ In y l /\ R x y. 23 | 24 | Definition is_restricted := forall x y, R x y -> In x l /\ In y l. 25 | 26 | Lemma incl_restriction : restriction << R. 27 | 28 | Proof. unfold restriction. repeat intro. tauto. Qed. 29 | 30 | Lemma restriction_dec : eq_dec A -> rel_dec R -> rel_dec restriction. 31 | 32 | Proof. 33 | unfold restriction. intros H H0 x y. destruct (H0 x y). 34 | destruct (In_dec H x l). destruct (In_dec H y l). 35 | constructor. tauto. constructor 2. tauto. constructor 2. tauto. 36 | constructor 2. tauto. 37 | Qed. 38 | 39 | Lemma restriction_midex : 40 | eq_midex A -> rel_midex R -> rel_midex restriction. 41 | 42 | Proof. 43 | unfold restriction. do 4 intro. destruct (H0 x y). 44 | destruct (In_midex H x l). destruct (In_midex H y l). 45 | constructor. tauto. constructor 2. tauto. constructor 2. tauto. 46 | constructor 2. tauto. 47 | Qed. 48 | 49 | End restriction. 50 | 51 | Lemma restriction_monotonic : forall (R' R'' : relation A) l, 52 | R' << R'' -> restriction R' l << restriction R'' l. 53 | 54 | Proof. unfold restriction. repeat intro. pose (H x y). tauto. Qed. 55 | 56 | Lemma restricted_restriction : forall R l, is_restricted (restriction R l) l. 57 | 58 | Proof. unfold restriction, is_restricted. tauto. Qed. 59 | 60 | Lemma restricted_clos_trans : forall R l, 61 | is_restricted R l -> is_restricted (clos_trans R) l. 62 | 63 | Proof. unfold is_restricted. intros. induction H0. apply H. hyp. tauto. Qed. 64 | 65 | Lemma clos_trans_restricted_pair : forall R x y, 66 | is_restricted R (x::y::nil) -> clos_trans R x y -> R x y. 67 | 68 | Proof. 69 | intros. induction H0. hyp. 70 | destruct (restricted_clos_trans H H0_). 71 | destruct H1. rewrite H1 in H. rewrite H1. tauto. 72 | destruct H1. rewrite H1 in H. rewrite H1. tauto. 73 | contr. 74 | Qed. 75 | 76 | Lemma clos_trans_restricted_pair_bis : forall R x y, 77 | is_restricted R (x::y::nil) -> clos_trans R y x -> R y x. 78 | 79 | Proof. 80 | intros. induction H0. hyp. 81 | destruct (restricted_clos_trans H H0_). 82 | destruct H1. rewrite H1 in H. rewrite H1. tauto. 83 | destruct H1. rewrite H1 in H. rewrite H1. tauto. 84 | contr. 85 | Qed. 86 | 87 | Lemma clos_trans_restriction : forall (R : relation A) x y, 88 | R x y -> clos_trans (restriction R (x :: y :: nil)) x y. 89 | 90 | Proof. intros. constructor. unfold restriction. simpl. tauto. Qed. 91 | 92 | End On_Rel_Gen. 93 | -------------------------------------------------------------------------------- /Util/Relation/SCC_dec.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Leo Ducas, 2007-08-06 6 | 7 | We give a way to decide the SCC relation using the adjacency matrix. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import GDomainBij AdjMat RelSub ListExtras RelUtil 13 | LogicUtil BoundNat SCC ListNodup. 14 | 15 | Record SCC_dec_hyps : Type := mkSCC_dec_hyps { 16 | hyp_A : Type; 17 | hyp_eq_dec : forall x y : hyp_A, {x=y} + {~x=y}; 18 | hyp_Dom : list hyp_A; 19 | hyp_R : relation hyp_A; 20 | hyp_restriction : is_restricted hyp_R hyp_Dom; 21 | hyp_nodup: nodup hyp_Dom; 22 | hyp_R_dec : forall x y, {hyp_R x y} + {~hyp_R x y} 23 | }. 24 | 25 | Section SCC_effectif. 26 | 27 | Variable hyps : SCC_dec_hyps. 28 | 29 | Notation A := (hyp_A hyps). 30 | Notation eq_dec := (hyp_eq_dec hyps). 31 | Notation Dom := (hyp_Dom hyps). 32 | Notation R := (hyp_R hyps). 33 | Notation restriction := (hyp_restriction hyps). 34 | Notation Dom_nodup := (hyp_nodup hyps). 35 | Notation R_dec := (hyp_R_dec hyps). 36 | Notation dim := (length Dom). 37 | 38 | Definition SCC_mat_effective := 39 | let M := MoG dim (rel_on_nat Dom R) (rel_on_nat_dec Dom R R_dec) in 40 | SCC_mat M. 41 | 42 | (* The Matrix M is kept as an argument, so main file may only compute 43 | it once *) 44 | 45 | Definition SCC_effective M (H : M = SCC_mat_effective) x y := 46 | rel_on_dom eq_dec Dom (GoM M) x y. 47 | 48 | Lemma SCC_effective_exact : forall M (H : M = SCC_mat_effective) x y, 49 | SCC R x y <-> SCC_effective H x y. 50 | 51 | Proof. 52 | split; intros; unfold SCC_effective in *; unfold SCC_mat_effective in *; 53 | rewrite <- (dom_change_SCC eq_dec restriction Dom_nodup x y) in *; 54 | unfold rel_on_dom in *; 55 | destruct (find_first (eq x) (eq_dec x) Dom); auto with *; 56 | destruct (find_first (eq y) (eq_dec y) Dom); auto with *; 57 | subst; 58 | rewrite GoM_SCC in *. 59 | assert (rel_on_nat Dom R 60 | << GoM (MoG dim (rel_on_nat Dom R) (rel_on_nat_dec Dom R R_dec))). 61 | unfold inclusion; intros; 62 | rewrite GoM_MoG; intros. trivial. 63 | ded (rel_on_nat_is_restricted _ _ _ _ H1). 64 | do 2 rewrite <- In_nats_decr_lt in H2; trivial. 65 | ded (SCC_incl H); auto. 66 | assert (GoM (MoG dim (rel_on_nat Dom R) (rel_on_nat_dec Dom R R_dec)) 67 | << rel_on_nat Dom R). 68 | unfold inclusion; intros. 69 | rewrite GoM_MoG in H; intros; auto. 70 | ded (rel_on_nat_is_restricted _ _ _ _ H1). 71 | do 2 rewrite <- In_nats_decr_lt in H2; trivial. ded (SCC_incl H). auto. 72 | Qed. 73 | 74 | Lemma SCC_effective_dec : forall M (H : M = SCC_mat_effective) x y, 75 | {SCC_effective H x y} + {~SCC_effective H x y}. 76 | 77 | Proof. 78 | intros. unfold SCC_effective. subst. 79 | unfold rel_on_dom. 80 | destruct (find_first (eq x) (eq_dec x)); auto with *. 81 | destruct (find_first (eq y) (eq_dec y)); auto with *. 82 | apply GoM_dec. 83 | Defined. 84 | 85 | End SCC_effectif. 86 | -------------------------------------------------------------------------------- /Coccinelle/list_extensions/equiv_list.v: -------------------------------------------------------------------------------- 1 | (**************************************************************************) 2 | (* * *) 3 | (* _ * The Coccinelle Library / Evelyne Contejean *) 4 | (* * CNRS-LRI-Universite Paris Sud *) 5 | (* -/@|@\- * A3PAT Project *) 6 | (* -@ | @- * *) 7 | (* -\@|@/- * This file is distributed under the terms of the *) 8 | (* -v- * CeCILL-C licence *) 9 | (* * *) 10 | (**************************************************************************) 11 | 12 | From Stdlib Require Import List Relations Setoid. 13 | 14 | Lemma equiv_in_list : 15 | forall (A : Type) (R : relation A) s ll, (forall s t, In (s,t) ll -> R s t) -> 16 | In s (map (fun st => fst st) ll) -> 17 | exists t, exists ll1, exists ll2, R s t /\ ll = ll1 ++ (s,t) :: ll2. 18 | Proof. 19 | intros A R s ll E_ll s_in_ll1; rewrite in_map_iff in s_in_ll1; 20 | destruct s_in_ll1 as [[s' t] [s'_eq_s st_in_ll]]; simpl in s'_eq_s; subst s'; 21 | destruct (In_split _ _ st_in_ll) as [ll1 [ll2 H]]; 22 | exists t; exists ll1; exists ll2; split; trivial. 23 | apply E_ll; trivial. 24 | Qed. 25 | 26 | Lemma equiv_in_list_2 : 27 | forall (A : Type) (R : relation A) s ll, (forall s t, In (s,t) ll -> R s t) -> 28 | In s (map (fun st => snd st) ll) -> 29 | exists t, exists ll1, exists ll2, R t s /\ ll = ll1 ++ (t,s) :: ll2. 30 | Proof. 31 | intros A R s ll E_ll s_in_ll2; rewrite in_map_iff in s_in_ll2; 32 | destruct s_in_ll2 as [[t s'] [s'_eq_s ts_in_ll]]; simpl in s'_eq_s; subst s'; 33 | destruct (In_split _ _ ts_in_ll) as [ll1 [ll2 H]]; 34 | exists t; exists ll1; exists ll2; split; trivial. 35 | apply E_ll; trivial. 36 | Qed. 37 | 38 | Lemma equiv_swap : 39 | forall (A : Type) (R : relation A), 40 | forall ll, (forall s t, In (s,t) ll -> R s t) -> 41 | exists ll', (forall s t, In (s,t) ll' -> R t s) /\ 42 | map (fun st => fst st) ll = map (fun st => snd st) ll' /\ 43 | map (fun st => snd st) ll = map (fun st => fst st) ll' /\ 44 | (forall s t, In (s,t) ll <-> In (t,s) ll'). 45 | Proof. 46 | intros A R ll; induction ll as [ | [s t] ll]; intro E_ll. 47 | exists (@nil (A * A)); repeat split; trivial; contradiction. 48 | destruct IHll as [ll' [E_ll' [H1 [H2 H3]]]]. 49 | intros; apply E_ll; right; trivial. 50 | exists ((t,s) :: ll'); repeat split. 51 | intros t' s' [t's'_eq_ts | t's'_in_ll']. 52 | injection t's'_eq_ts; intros; subst; apply E_ll; left; trivial. 53 | apply E_ll'; trivial. 54 | simpl; rewrite H1; trivial. 55 | simpl; rewrite H2; trivial. 56 | intros [st_eq_st | st_in_ll]. 57 | injection st_eq_st; intros; subst; left; trivial. 58 | right; rewrite <- H3; trivial. 59 | intros [ts_eq_ts | ts_in_ll]. 60 | injection ts_eq_ts; intros; subst; left; trivial. 61 | right; rewrite H3; trivial. 62 | Qed. 63 | 64 | 65 | 66 | 67 | 68 | 69 | 70 | 71 | 72 | -------------------------------------------------------------------------------- /TODO: -------------------------------------------------------------------------------- 1 | 2 | - move VecEq into VecOrd 3 | 4 | - replace LexOrder by Lexico? 5 | 6 | - define a feq generalizing the ones in ListUtil and FSetUtil 7 | 8 | - in RelUtil, replace # by * and ! by + 9 | 10 | - replace Add [Parametric] Morphism by Proper 11 | 12 | - remove Implicit in Implicit Arguments 13 | 14 | - define default_mi 15 | 16 | - factorize the part of the proof that is similar in 17 | AFilterPerm.filter_cont_closed and 18 | AFilterPerm.red_incl_filter_red_rc. 19 | 20 | - in BoundNat, replace bnat by nat_lt 21 | 22 | - replace decidability lemmas by boolean functions 23 | in polynomial and matrix interpretations 24 | 25 | - for polynomial monotony, check that, all coefficients are 26 | non-negative and that, for all i, there is a monomial [k_1,..,k_n] 27 | which coefficient is positive and such that k_i is positive too 28 | 29 | - semantic labelling with booleans 30 | 31 | - FC2 32 | 33 | - usable rules 34 | 35 | - precedence termination? (special case of RPO) 36 | 37 | - unify Arctic and Arctic-Below-Zero methods 38 | 39 | - factorize AMatrixInt and ABigMatrixInt? 40 | 41 | - define hd_red_mod from hd_red_Mod 42 | 43 | - only defined symbols need to be marked 44 | 45 | - replace plus by tail recursive plus 46 | 47 | - make problem statement (including signature) explicit at every proof step? 48 | 49 | - replace defs and lemmas when they are added in coq std lib 50 | 51 | - replace unary integers by binary integers (big nats)? 52 | 53 | Done: 54 | ----- 55 | 56 | - in AFilterPerm, prove that (map (@val (arity f)) (pi f) = raw_pi f) 57 | and define permut on raw_pi instead of pi. 58 | 59 | - extend filter to red in case of permutations only 60 | 61 | - remove MPO top level hypotheses -> make a functor 62 | 63 | - extend Term/WithArity/AInfSeq.v to arbitrary relations (not only on 64 | terms) and rename it in to Util/Relation/InfSeq.v 65 | 66 | - uniformize the various notions of bounded nat in NatUtil, and use a 67 | record definition for defining projections at the same time 68 | 69 | - remove implicit arguments from tactic absolute_finite, 70 | fin_monotone_succ, prove_cc, etc.: Fs can be inferred from Fs_ok 71 | 72 | - add improvement to DP pairs by ignoring pairs where rhs is a 73 | subterm of lhs 74 | 75 | - add boolean function for rpo 76 | 77 | - replace showArctic[BZ]IntOk by boolean test 78 | 79 | - use Qed in beq_*_ok 80 | 81 | - replace rev by rev' 82 | 83 | - AF: extend by allowing collapsing argument filterings 84 | 85 | - prove that red R is finitely branching 86 | 87 | - remove from RelExtras the things already defined in RelUtil 88 | 89 | - use RelUtil notations in RelExtras 90 | 91 | - extend WF_sred_of_WF_red to red_mod 92 | 93 | - remove the use of beq_dec and beqtac 94 | 95 | - define eq_term_dec from beq_term 96 | 97 | - absurd_hyp is OBSOLETE: use contradict instead. 98 | 99 | - Eqdep uses an axiom: it is still used in ATerm and ACap (grep Eqdep) 100 | and in the tactic Funeqtac. is it possible to get rid of it by using 101 | Eqdep_dec only? yes 102 | 103 | - DP: improve by introducing marked symbols 104 | 105 | - move ListExt into ListUtil 106 | -------------------------------------------------------------------------------- /Term/WithArity/ATrsNorm.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2008-08-08 6 | 7 | 1) canonical representation of a rule l -> r: 8 | 9 | WE ASSUME: incl (vars r) (vars l). 10 | 11 | given the list of variables of l, xs = ATerm.vars l, 12 | a variable of l or r is renamed as its position in xs 13 | 14 | 2) canonical representation of a list of rules R: 15 | 16 | the rules of R are put in their unique representation 17 | and R is sorted wrt the following ordering 18 | 19 | (l1,r1) >rule (l2,r2) if (l1,r1) (>term)lex (l2,r2) 20 | 21 | f(t1..tn) >term g(u1..up) if 22 | (f,(t1..tn)) (>symb,(>term)lex)lex (g,(u1..up)) 23 | 24 | f(t1..tn) >term x 25 | 26 | x >term y if x >nat y 27 | 28 | assuming a total ordering >symb on symbols 29 | *) 30 | 31 | Set Implicit Arguments. 32 | 33 | From Stdlib Require Import List Compare_dec. 34 | From CoLoR Require Import ATrs ListDec ListSort NatUtil VecUtil. 35 | 36 | Section S. 37 | 38 | Variable Sig : Signature. 39 | 40 | Notation term := (term Sig). Notation terms := (vector term). 41 | 42 | (***********************************************************************) 43 | (** canonical representation of a rule *) 44 | 45 | Section norm. 46 | 47 | Variable xs : variables. 48 | 49 | Fixpoint norm (t : term) : term := 50 | match t with 51 | | Var x => 52 | match position beq_nat x xs with 53 | | Some y => Var y 54 | | None => Var (length xs) 55 | end 56 | | Fun f ts => Fun f (Vmap norm ts) 57 | end. 58 | 59 | End norm. 60 | 61 | (* ASSUME: incl (vars r) (vars l) *) 62 | Definition mkNormRule l r := 63 | let xs := ATerm.vars l in mkRule (norm xs l) (norm xs r). 64 | 65 | (***********************************************************************) 66 | (** canonical representation of a list of rules *) 67 | 68 | Section term_ordering. 69 | 70 | Variable symb_cmp : Sig -> Sig -> comparison. 71 | 72 | Fixpoint cmp (t u : term) : comparison := 73 | match t, u with 74 | | Var x, Var y => Nat.compare x y 75 | | Var x, _ => Lt 76 | | _, Var x => Gt 77 | | Fun f ts, Fun g us => 78 | match symb_cmp f g with 79 | | Eq => 80 | let fix cmp_terms n (ts : terms n) p (us : terms p) 81 | : comparison := 82 | match ts, us with 83 | | Vnil, Vnil => Eq 84 | | Vnil, _ => Lt 85 | | _, Vnil => Gt 86 | | Vcons t ts, Vcons u us => 87 | match cmp t u with 88 | | Eq => cmp_terms _ ts _ us 89 | | c => c 90 | end 91 | end 92 | in cmp_terms (arity f) ts (arity g) us 93 | | c => c 94 | end 95 | end. 96 | 97 | (* ASSUME: rules are in normal form with mkNormRule *) 98 | Definition mkNormRules := sort cmp. 99 | 100 | End term_ordering. 101 | 102 | End S. 103 | -------------------------------------------------------------------------------- /Util/Option/OptUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2009-03-27 6 | 7 | Utility results about the option/exception type. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import Bool Program. 13 | From CoLoR Require Import ListUtil LogicUtil. 14 | Import ListNotations. 15 | 16 | Lemma Some_eq : forall A (x y : A), Some x = Some y -> x = y. 17 | 18 | Proof. intros. inversion H. auto. Qed. 19 | 20 | Lemma option_dec A (el : option A) : el = None \/ exists w, el = Some w. 21 | 22 | Proof. 23 | intros. destruct el. 24 | right; exists a; trivial. 25 | left; trivial. 26 | Qed. 27 | 28 | Section dec. 29 | 30 | Variables (A : Type) (eq_dec : forall x y : A, {x=y}+{~x=y}). 31 | 32 | Lemma eq_opt_dec : forall x y : option A, {x=y}+{~x=y}. 33 | 34 | Proof. decide equality. Qed. 35 | 36 | End dec. 37 | 38 | Section OptUtil. 39 | 40 | Variable A : Type. 41 | Variable P : A -> Prop. 42 | Variable f : forall a : A, option (P a). 43 | 44 | Definition opt_to_bool (P : Prop) (e : option P) := 45 | match e with 46 | | Some _ => true 47 | | None => false 48 | end. 49 | 50 | Definition pred_opt_to_bool x := opt_to_bool (f x). 51 | 52 | Program Definition lforall_opt (l : list A) : 53 | option (lforall (fun x : A => exists p, f x = Some p) l) := 54 | match forallb pred_opt_to_bool l with 55 | | true => Some _ 56 | | false => None 57 | end. 58 | 59 | Next Obligation. 60 | Proof with auto. 61 | apply forallb_imp_lforall with (fun x => opt_to_bool (f x))... 62 | intros. destruct (f x). exists p... discr. 63 | Qed. 64 | 65 | Program Definition partition_opt (l : list A) : 66 | { ll : list A * list A | 67 | let (l1, l2) := ll in 68 | lforall P l1 /\ 69 | (lforall (fun x => f x = error) l2) /\ 70 | (forall x, In x l -> In x l1 \/ In x l2) 71 | } := partition pred_opt_to_bool l. 72 | 73 | Next Obligation. 74 | Proof with auto. 75 | assert (left := fun x => partition_left pred_opt_to_bool x l). 76 | assert (right := fun x => partition_right pred_opt_to_bool x l). 77 | assert (complete := fun x => partition_complete pred_opt_to_bool x l). 78 | destruct (partition pred_opt_to_bool l). 79 | unfold pred_opt_to_bool in *. 80 | repeat split... 81 | apply lforall_intro. intros x xl. 82 | assert (ll := left x xl). destruct (f x); try discr... 83 | apply lforall_intro. intros x xl. 84 | assert (rr := right x xl). destruct (f x); try discr... 85 | Qed. 86 | 87 | End OptUtil. 88 | 89 | Section Map_option. 90 | 91 | Variable A B : Type. 92 | Variable f : A -> option B. 93 | 94 | Program Fixpoint map_opt (l : list A) : option (list B) := 95 | match l with 96 | | [] => Some [] 97 | | cons x xs => 98 | match f x with 99 | | None => None 100 | | Some v => 101 | match map_opt xs with 102 | | None => None 103 | | Some vs => Some (v :: vs) 104 | end 105 | end 106 | end. 107 | 108 | End Map_option. 109 | 110 | Arguments opt_to_bool [P] _. 111 | -------------------------------------------------------------------------------- /Coccinelle/examples/cime_trace/term_extension.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import List Bool. 2 | From CoLoR Require Import TransClosure. 3 | From CoLoR Require term terminaison. 4 | 5 | Module Make(M:term_spec.Term). 6 | Import M. 7 | 8 | 9 | Ltac find_replacement l := 10 | match l with 11 | | nil => constr:(@nil term) 12 | | ?x::?l => 13 | let l':=find_replacement l in 14 | let e := 15 | match goal with 16 | | H: refl_trans_clos _ ?x' x |- _ => 17 | constr:(x') 18 | | |- _ => constr:(x) 19 | end 20 | in 21 | constr:(e::l') 22 | end. 23 | 24 | 25 | Ltac is_closed_list is_closed l := 26 | match l with 27 | | nil => idtac 28 | | ?t::?l => 29 | is_closed t; 30 | is_closed_list is_closed l 31 | end 32 | . 33 | 34 | Ltac is_closed t := 35 | match t with 36 | | Term _ ?l => is_closed_list ltac:(is_closed l) 37 | | Var _ => idtac 38 | | _ => 39 | (match goal with 40 | | H:_ |- _ => 41 | match t with 42 | | H => idtac 43 | end 44 | end) 45 | end 46 | . 47 | 48 | Ltac star_refl' := 49 | repeat (match goal with 50 | | H:_ |- _ => clear H|| revert H 51 | end);intros; 52 | match goal with 53 | | |- refl_trans_clos _ ?t2 ?t1 => 54 | (is_closed t1 ; econstructor 1) || (is_closed t2; econstructor 1) 55 | end. 56 | 57 | 58 | Definition eq_term_term_bool t1_t2 t1'_t2' : bool := 59 | match t1_t2,t1'_t2' with 60 | | (t1,t2),(t1',t2') => andb (M.eq_bool t1 t1') (M.eq_bool t2 t2') 61 | end. 62 | 63 | Lemma eq_term_term_bool_ok : forall t1_t2 t1'_t2', 64 | if eq_term_term_bool t1_t2 t1'_t2' then t1_t2 = t1'_t2' else t1_t2 <> t1'_t2'. 65 | Proof. 66 | intros [t1 t2] [t1' t2']. 67 | simpl. 68 | generalize (M.eq_bool_ok t1 t1'). 69 | case (eq_bool t1 t1'). 70 | generalize (M.eq_bool_ok t2 t2'). 71 | case (eq_bool t2 t2'). 72 | intros;subst;reflexivity. 73 | intros t2_neq_t2' _ abs;elim t2_neq_t2';injection abs;intros;assumption. 74 | intros t1_neq_t1' abs;elim t1_neq_t1';injection abs;intros;assumption. 75 | Defined. 76 | 77 | 78 | 79 | Lemma term_eq_bool_equiv : 80 | forall f g : M.term, f = g <-> M.eq_bool f g = true. 81 | Proof. 82 | intros f g. 83 | generalize (M.eq_bool_ok f g). 84 | case (eq_bool f g). 85 | tauto. 86 | intuition discriminate. 87 | Qed. 88 | 89 | 90 | 91 | End Make. 92 | 93 | 94 | Module IntVars. 95 | Definition A := nat. 96 | Fixpoint eq_bool (n1 n2:nat) : bool := 97 | match n1,n2 with 98 | | 0,0 => true 99 | | S n1,S n2 => eq_bool n1 n2 100 | | _,_ => false 101 | end. 102 | 103 | Lemma eq_bool_ok : forall n1 n2, if eq_bool n1 n2 then n1=n2 else n1<>n2. 104 | Proof. 105 | induction n1 as [|n1];intros [|n2];simpl in *. 106 | 107 | exact (eq_refl _). 108 | intro abs;discriminate abs. 109 | intro abs;discriminate abs. 110 | generalize (IHn1 n2). 111 | case (eq_bool n1 n2). 112 | intro Heq;f_equal;assumption. 113 | intros Hneq abs;elim Hneq;injection abs;intro H;exact H. 114 | Defined. 115 | End IntVars. 116 | -------------------------------------------------------------------------------- /Util/Function/FunUtil.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2014-12-11 6 | 7 | useful definitions and lemmas on functions 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil. 13 | 14 | (****************************************************************************) 15 | (** Some properties of functions. *) 16 | 17 | Section functions. 18 | 19 | Variables (A B : Type) (f : A -> B). 20 | 21 | Definition injective := forall x y, f x = f y -> x = y. 22 | 23 | Definition surjective := forall y, exists x, y = f x. 24 | 25 | Definition bijective := injective /\ surjective. 26 | 27 | End functions. 28 | 29 | (****************************************************************************) 30 | (** Function composition. *) 31 | 32 | Section comp. 33 | 34 | Variables (A B C : Type) (f : B -> C) (g : A -> B). 35 | 36 | Definition comp x := f (g x). 37 | 38 | Lemma inj_comp : injective f -> injective g -> injective comp. 39 | 40 | Proof. 41 | intros f_inj g_inj x y e. apply f_inj in e. apply g_inj in e. hyp. 42 | Qed. 43 | 44 | Lemma inj_comp_elim : injective comp -> injective g. 45 | 46 | Proof. 47 | intros comp_inj x y e. apply comp_inj. unfold comp. rewrite e. refl. 48 | Qed. 49 | 50 | Lemma surj_comp : surjective f -> surjective g -> surjective comp. 51 | 52 | Proof. 53 | intros f_surj g_surj z. destruct (f_surj z) as [y hy]. 54 | destruct (g_surj y) as [x hx]. ex x. subst. refl. 55 | Qed. 56 | 57 | Lemma surj_comp_elim : surjective comp -> surjective f. 58 | 59 | Proof. 60 | intros comp_surj y. destruct (comp_surj y) as [x e]. subst. ex (g x). refl. 61 | Qed. 62 | 63 | Lemma bij_comp : bijective f -> bijective g -> bijective comp. 64 | 65 | Proof. 66 | intros f_bij g_bij. split. apply inj_comp; fo. apply surj_comp; fo. 67 | Qed. 68 | 69 | End comp. 70 | 71 | Infix "o" := comp (at level 70). 72 | 73 | (****************************************************************************) 74 | (** Inverse of a surjective function, using Hilbert's epsilon operator. *) 75 | 76 | From CoLoR Require Import EpsilonUtil. 77 | 78 | Section inverse. 79 | 80 | Variables (A B : Type) (f : A -> B) (f_surj : surjective f). 81 | 82 | Definition inverse : B -> A. 83 | 84 | Proof. intro y. destruct (cid (f_surj y)) as [x _]. exact x. Defined. 85 | 86 | Lemma inverse_eq y : f (inverse y) = y. 87 | 88 | Proof. unfold inverse. destruct (cid (f_surj y)) as [x e]. auto. Qed. 89 | 90 | Lemma inj_inverse : injective inverse. 91 | 92 | Proof. 93 | intros x y e. apply (f_equal f) in e. rewrite !inverse_eq in e. hyp. 94 | Qed. 95 | 96 | End inverse. 97 | 98 | (****************************************************************************) 99 | (** Tactics for injectivity, surjective and bijectivity. *) 100 | 101 | Ltac inj := 102 | match goal with 103 | | |- injective (_ o _) => apply inj_comp; inj 104 | | |- injective (inverse _) => apply inj_inverse; inj 105 | | |- injective _ => auto 106 | end. 107 | 108 | Ltac surj := 109 | match goal with 110 | | |- surjective (_ o _) => apply surj_comp; surj 111 | | |- surjective _ => auto 112 | end. 113 | 114 | Ltac bij := 115 | match goal with 116 | | |- bijective (_ o _) => apply bij_comp; bij 117 | | |- bijective _ => hyp || (split; [inj | surj]) 118 | end. 119 | -------------------------------------------------------------------------------- /Term/Varyadic/VTrs.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-02-17 6 | 7 | rewriting 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Export VContext VSubstitution LogicUtil. 13 | From CoLoR Require Import RelUtil. 14 | 15 | (***********************************************************************) 16 | (** definition *) 17 | 18 | Section def. 19 | 20 | Variable Sig : Signature. 21 | 22 | Notation term := (term Sig). 23 | 24 | Record rule : Type := mkRule { lhs : term; rhs : term }. 25 | 26 | Definition rules := list rule. 27 | 28 | Definition red R t1 t2 := exists l, exists r, exists c, exists s, 29 | In (mkRule l r) R /\ t1 = fill c (sub s l) /\ t2 = fill c (sub s r). 30 | 31 | Definition hd_red R t1 t2 := exists l, exists r, exists s, 32 | In (mkRule l r) R /\ t1 = sub s l /\ t2 = sub s r. 33 | 34 | Definition int_red R t1 t2 := exists l, exists r, exists c, exists s, 35 | c <> Hole 36 | /\ In (mkRule l r) R /\ t1 = fill c (sub s l) /\ t2 = fill c (sub s r). 37 | 38 | Definition red_mod E R := red E # @ red R. 39 | 40 | End def. 41 | 42 | (***********************************************************************) 43 | (** tactics *) 44 | 45 | Ltac redtac := repeat 46 | match goal with 47 | | H : red ?R ?t ?u |- _ => 48 | let l := fresh "l" in let r := fresh "r" in let c := fresh "c" in 49 | let s := fresh "s" in let h1 := fresh in 50 | (unfold red in H; destruct H as [l]; destruct H as [r]; destruct H as [c]; 51 | destruct H as [s]; destruct H as [H h1]; destruct h1) 52 | | H : transp (red _) _ _ |- _ => unfold transp in H; redtac 53 | | H : hd_red ?R ?t ?u |- _ => 54 | let l := fresh "l" in let r := fresh "r" in 55 | let s := fresh "s" in let h1 := fresh in 56 | (unfold hd_red in H; destruct H as [l]; destruct H as [r]; 57 | destruct H as [s]; destruct H as [H h1]; destruct h1) 58 | | H : transp (hd_red _) _ _ |- _ => unfold transp in H; redtac 59 | | H : int_red ?R ?t ?u |- _ => 60 | let l := fresh "l" in let r := fresh "r" in let c := fresh "c" in 61 | let s := fresh "s" in let h1 := fresh in let h2 := fresh in 62 | (unfold int_red in H; destruct H as [l]; destruct H as [r]; 63 | destruct H as [c]; destruct H as [s]; destruct H as [H h1]; 64 | destruct h1 as [h1 h2]; destruct h2) 65 | | H : transp (int_red _) _ _ |- _ => 66 | unfold transp in H; redtac 67 | end. 68 | 69 | (***********************************************************************) 70 | (** properties *) 71 | 72 | Section S. 73 | 74 | Variable Sig : Signature. 75 | 76 | Notation rule := (rule Sig). 77 | 78 | Variable R : list rule. 79 | 80 | Lemma red_rule : forall l r c s, 81 | In (mkRule l r) R -> red R (fill c (sub s l)) (fill c (sub s r)). 82 | 83 | Proof. 84 | intros. unfold red. exists l. exists r. exists c. exists s. auto. 85 | Qed. 86 | 87 | Lemma red_rule_top : forall l r s, 88 | In (mkRule l r) R -> red R (sub s l) (sub s r). 89 | 90 | Proof. 91 | intros. unfold red. exists l. exists r. exists (@Hole Sig). exists s. auto. 92 | Qed. 93 | 94 | Lemma red_fill : forall c t u, red R t u -> red R (fill c t) (fill c u). 95 | 96 | Proof. 97 | intros. redtac. unfold red. 98 | exists l. exists r. exists (comp c c0). exists s. split. hyp. 99 | subst t. subst u. do 2 rewrite fill_fill. auto. 100 | Qed. 101 | 102 | End S. 103 | -------------------------------------------------------------------------------- /Coccinelle/examples/cime_trace/make_or.ml: -------------------------------------------------------------------------------- 1 | open Format 2 | 3 | let rec pr_ais fmt i = 4 | if i>0 5 | then 6 | fprintf fmt "%a A%d " pr_ais (i-1) i 7 | 8 | 9 | let rec pr_or_ais fmt i = 10 | if i>1 11 | then 12 | fprintf fmt "%a \\/ A%d " pr_or_ais (i-1) i 13 | else fprintf fmt "A1 " 14 | 15 | 16 | let pr_ori_constr fmt i j = 17 | fprintf fmt "| or%d_%d : A%d -> or%d %a\n" 18 | i j j i pr_ais i 19 | 20 | let pr_ori_ind fmt i = 21 | fprintf fmt 22 | "@[Inductive or%d (%a:Prop) : Prop :=@." 23 | i pr_ais i; 24 | for j=1 to i do 25 | pr_ori_constr fmt i j; 26 | done; 27 | fprintf fmt ".@]@.@." 28 | 29 | 30 | (* let pr_ori_lemma1 fmt i j = *) 31 | (* for k = 1 to j -1 do *) 32 | (* fprintf fmt "right. " *) 33 | (* done; *) 34 | (* if (i<>j) then fprintf fmt "left. " *) 35 | 36 | let pr_ori_lemma1 fmt i = 37 | if i > 3 38 | then 39 | begin 40 | fprintf fmt "rewrite <- or%d_equiv.@." (i-1); 41 | fprintf fmt "destruct H.@."; 42 | for j=1 to (i-2) do 43 | fprintf fmt "constructor %d. exact H.@." j 44 | done; 45 | fprintf fmt "constructor %d. left. exact H.@." (i-1); 46 | fprintf fmt "constructor %d. right. exact H.@." (i-1); 47 | end 48 | else 49 | begin 50 | fprintf fmt "destruct H.@."; 51 | fprintf fmt "left. assumption.@."; 52 | fprintf fmt "right. left. assumption.@."; 53 | fprintf fmt "right. right. assumption.@." 54 | end 55 | 56 | let pr_ori_lemma2 fmt i = 57 | if i = 3 58 | then 59 | begin 60 | for j = 1 to (i-1) do 61 | fprintf fmt "destruct H. constructor %d. exact H.@." j 62 | done; 63 | fprintf fmt "constructor %d. exact H.@." i 64 | end 65 | else 66 | begin 67 | fprintf fmt "destruct H.@."; 68 | fprintf fmt "constructor 1. exact H.@."; 69 | fprintf fmt "rewrite <- or%d_equiv in H.@." (i-1); 70 | fprintf fmt "destruct H.@."; 71 | for j = 2 to i do 72 | fprintf fmt "constructor %d. exact H.@." j; 73 | done; 74 | 75 | end 76 | 77 | let pr_ori_lemma fmt i = 78 | fprintf fmt "Lemma or%d_equiv: forall %a, or%d %a <-> %a.@." 79 | i pr_ais i i pr_ais i pr_or_ais i; 80 | fprintf fmt "@[Proof. "; 81 | fprintf fmt "intros.@.split.@.@."; 82 | fprintf fmt "intros H.@."; 83 | (* fprintf fmt "repeat (first[left;assumption|right]);assumption.@.@."; *) 84 | (* for j = 1 to i do pr_ori_lemma1 fmt i j; fprintf fmt "assumption.@."; done; *) 85 | pr_ori_lemma1 fmt i; 86 | fprintf fmt "intros H.@."; 87 | (* fprintf fmt "repeat (destruct H;[constructor assumption|idtac]);constructor assumption.@]@.Qed.@.@.@." *) 88 | pr_ori_lemma2 fmt i; 89 | fprintf fmt "@]@.Qed.@.@.@." 90 | 91 | 92 | let pr_ori_morph fmt i = 93 | fprintf fmt "Add Morphism or%d : or%d_morph.@." i i; 94 | fprintf fmt "Proof.@."; 95 | fprintf fmt "intros.@."; 96 | fprintf fmt "do 2 rewrite or%d_equiv.@." i; 97 | fprintf fmt "repeat match goal with@."; 98 | fprintf fmt "H: _ <-> _ |- _ => rewrite H;clear H@."; 99 | fprintf fmt "end.@."; 100 | fprintf fmt "reflexivity.@."; 101 | fprintf fmt "Qed.@." 102 | 103 | let pr_ori fmt i = 104 | pr_ori_ind fmt i; 105 | pr_ori_lemma fmt i; 106 | pr_ori_morph fmt i 107 | 108 | 109 | 110 | 111 | let _ = 112 | let f = open_out "or_ext_generated.v" in 113 | let fmt = formatter_of_out_channel f in 114 | fprintf fmt "Require Import Setoid.@."; 115 | fprintf fmt "Set Implicit Arguments.@.@.@."; 116 | for i = 3 to 100 do pr_ori fmt i; done; 117 | close_out f 118 | -------------------------------------------------------------------------------- /Util/Multiset/MultisetCore.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2004-09-06 6 | 7 | This file provides a specification of finite multiset data-type along 8 | with specification of operations on multisets. 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | From CoLoR Require RelExtras. 14 | From Stdlib Require Import Relations Psatz. 15 | 16 | Declare Scope msets_scope. 17 | 18 | Module Type MultisetCore. 19 | 20 | Declare Module Export Sid : RelExtras.Eqset_dec. 21 | 22 | Parameter Multiset : Type. 23 | 24 | Parameter mult : A -> Multiset -> nat. 25 | 26 | Parameter meq : relation Multiset. 27 | 28 | Parameter empty : Multiset. 29 | 30 | Parameter singleton : A -> Multiset. 31 | 32 | Parameter union : Multiset -> Multiset -> Multiset. 33 | 34 | Parameter intersection : Multiset -> Multiset -> Multiset. 35 | 36 | Parameter diff : Multiset -> Multiset -> Multiset. 37 | 38 | Notation "X =mul= Y" := (meq X Y) (at level 70) : msets_scope. 39 | Notation "X <>mul Y" := (~meq X Y) (at level 50) : msets_scope. 40 | Notation "{{ x }}" := (singleton x) (at level 5) : msets_scope. 41 | Notation "X + Y" := (union X Y) : msets_scope. 42 | Notation "X - Y" := (diff X Y) : msets_scope. 43 | Notation "X # Y" := (intersection X Y) (at level 50, left associativity) 44 | : msets_scope. 45 | Notation "x / M" := (mult x M) : msets_scope. 46 | 47 | Delimit Scope msets_scope with msets. 48 | Open Scope msets_scope. 49 | Bind Scope msets_scope with Multiset. 50 | 51 | Section Specification. 52 | 53 | Variables (M N P : Multiset) (x y z : A) (n : nat). 54 | 55 | Parameter mult_eqA_compat: x =A= y -> x/M = y/M. 56 | 57 | Parameter meq_multeq: M =mul= N -> (forall x, x/M = x/N). 58 | 59 | Parameter multeq_meq: (forall x, x/M = x/N) -> M =mul= N. 60 | 61 | Parameter empty_mult: x/empty = 0. 62 | 63 | Parameter union_mult: x/(M+N) = (x/M + x/N)%nat. 64 | 65 | Parameter diff_mult: x/(M-N) = (x/M - x/N)%nat. 66 | 67 | Parameter intersection_mult: x/(M#N) = min (x/M) (x/N). 68 | 69 | Parameter singleton_mult_in: x =A= y -> x/{{y}} = 1. 70 | 71 | Parameter singleton_mult_notin: ~x =A= y -> x/{{y}} = 0. 72 | 73 | Parameter mset_ind_type: forall P : Multiset -> Type, 74 | P empty -> (forall M a, P M -> P (M + {{a}})) -> forall M, P M. 75 | 76 | End Specification. 77 | 78 | #[global] Hint Resolve mult_eqA_compat 79 | meq_multeq 80 | multeq_meq 81 | empty_mult 82 | union_mult 83 | diff_mult 84 | intersection_mult 85 | singleton_mult_in 86 | singleton_mult_notin : multisets. 87 | 88 | Global Hint Rewrite empty_mult 89 | union_mult 90 | diff_mult 91 | intersection_mult using trivial : multisets. 92 | 93 | End MultisetCore. 94 | 95 | Section Multiset_IntersectionAsDifference. 96 | 97 | Variables 98 | (A : Type) (Multiset : Type) (meq : relation Multiset) 99 | (mult : A -> Multiset -> nat) 100 | (diff : Multiset -> Multiset -> Multiset) 101 | (diff_mult : forall M N x, mult x (diff M N) = mult x M - mult x N). 102 | 103 | Definition inter_as_diff x y := diff x (diff x y). 104 | 105 | Lemma inter_as_diff_ok : forall M N x, 106 | mult x (inter_as_diff M N) = Nat.min (mult x M) (mult x N). 107 | 108 | Proof. 109 | intros M N x. unfold inter_as_diff. rewrite !diff_mult. lia. 110 | Qed. 111 | 112 | End Multiset_IntersectionAsDifference. 113 | -------------------------------------------------------------------------------- /Util/Nat/Log2.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Leo Ducas, 2007-08-06 6 | 7 | Definition of log2 (floor) and exp2, and some equalities 8 | *) 9 | 10 | From Stdlib Require Import Lia PeanoNat. 11 | From CoLoR Require Import LogicUtil NatCompat. 12 | 13 | Lemma div2_le_n : forall n, Nat.div2 n <= n. 14 | 15 | Proof. 16 | cut (forall n, Nat.div2 n <= n /\ Nat.div2 (S n) <= S n). 17 | intros H n; gen (H n); tauto. 18 | induction n; auto. 19 | inversion IHn; split; auto. 20 | simpl. apply le_n_S. auto with arith. 21 | Qed. 22 | 23 | Inductive log2_prop : nat -> nat -> Prop := 24 | | log2_prop_O : log2_prop 0 0 25 | | log2_prop_1 : log2_prop 1 0 26 | | log2_prop_p : forall p q, 27 | p <> 0 -> p <> 1 -> log2_prop (Nat.div2 p) q -> log2_prop p (S q). 28 | 29 | #[global] Hint Constructors log2_prop : core. 30 | 31 | Fixpoint log2_aux n count : nat := 32 | match count with 33 | | 0 => 0 34 | | S count' => 35 | match n with 36 | | 0 => 0 37 | | 1 => 0 38 | | _ => S (log2_aux (Nat.div2 n) count') 39 | end 40 | end. 41 | 42 | (* log2_aux is correct as soon as count >= n *) 43 | 44 | Lemma log2_aux_matches : forall count n, n <= count -> 45 | log2_prop n (log2_aux n count). 46 | 47 | Proof. 48 | induction count. 49 | intros. assert (n = 0). auto with arith. 50 | subst. simpl. auto. intros n Hn. 51 | destruct n. 52 | simpl; auto. 53 | destruct n. 54 | simpl; auto. 55 | simpl. 56 | apply log2_prop_p. 57 | intro; discr. 58 | intro H; discr. 59 | apply IHcount; eapply Nat.le_trans. 60 | eapply le_n_S; apply div2_le_n. 61 | apply le_S_n; hyp. 62 | Qed. 63 | 64 | Definition log2 n := log2_aux n n. 65 | 66 | Lemma log2_matches : forall n, log2_prop n (log2 n). 67 | 68 | Proof. unfold log2. intros. apply log2_aux_matches. auto with arith. Qed. 69 | 70 | Lemma log2_prop_func : forall p q, 71 | log2_prop p q -> forall q', log2_prop p q' -> q = q'. 72 | 73 | Proof. 74 | induction 1. 75 | inversion 1. 76 | auto. 77 | congruence. 78 | inversion 1. 79 | auto. 80 | congruence. 81 | inversion 1. 82 | congruence. 83 | congruence. 84 | subst. 85 | f_equal. 86 | eapply IHlog2_prop; eauto. 87 | Qed. 88 | 89 | Lemma log2_matches_log2_prop : forall n p, log2_prop n p -> p = log2 n. 90 | 91 | Proof. intros. eapply log2_prop_func. ehyp. apply log2_matches. Qed. 92 | 93 | Fixpoint exp2 n := 94 | match n with 95 | | O => 1 96 | | S i => 2 * exp2 i 97 | end. 98 | 99 | Lemma double_div2 : forall n, S (2 * Nat.div2 n) >= n. 100 | 101 | Proof. 102 | intro n. destruct (Nat.Even_or_Odd n). 103 | rewrite NatCompat.Even_double; [| exact H]; unfold Nat.double; lia. 104 | rewrite NatCompat.Odd_double; auto; unfold Nat.double; lia. 105 | Qed. 106 | 107 | Lemma exp2_pos : forall n, exp2 n >0. 108 | 109 | Proof. intro; induction n; simpl; auto with *. Qed. 110 | 111 | Lemma exp2_log2_prop : forall n p, log2_prop n p -> exp2 (S p) > n. 112 | 113 | Proof. 114 | cut(forall n n' p, n'<=n -> log2_prop n' p -> exp2 (S p) > n'). 115 | intros. eapply H. assert (n<=n). lia. ehyp. hyp. 116 | intro; induction n; intros. assert (n'=0). lia. subst. apply exp2_pos. 117 | inversion H; auto. 118 | subst; inversion H0. simpl; auto. 119 | subst. 120 | assert (exp2 (S q) > Nat.div2 (S n)). 121 | apply IHn. 122 | simpl; destruct n; auto with *. 123 | assert (Nat.div2 n <= n); auto with *. apply div2_le_n. 124 | hyp. 125 | change (2*exp2 (S q) > S n). 126 | gen(double_div2 (S n)); intro; lia. 127 | Qed. 128 | 129 | Lemma exp2_log2 : forall n, exp2(S (log2 n)) > n. 130 | 131 | Proof. intro; apply exp2_log2_prop. apply log2_matches. Qed. 132 | -------------------------------------------------------------------------------- /Util/Nat/NatCompat.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import Arith.PeanoNat. 2 | (***********************************************************************) 3 | (** Adapt to the removal of some Arith files in 8.19: 4 | unidirectional version better suited to implicit arguments. 5 | WARNING: we want these to shadow some eventual remnants in Arith, so until 6 | Coq 8.19, one needs to be careful to Require NatUtil **after** Coq.Arith 7 | or to qualify these with NatUtil. 8 | *) 9 | 10 | Lemma lt_S_n : forall n m : nat, S n < S m -> n < m. 11 | Proof. intros n m; now apply Nat.succ_lt_mono. Qed. 12 | 13 | Lemma lt_n_S : forall n m : nat, n < m -> S n < S m. 14 | Proof. intros n m; now apply Nat.succ_lt_mono. Qed. 15 | 16 | Lemma gt_le_S : forall n m : nat, m > n -> S n <= m. 17 | Proof. intros n m; now apply Nat.le_succ_l. Qed. 18 | 19 | Lemma le_lt_n_Sm : forall n m : nat, n <= m -> n < S m. 20 | Proof. intros n m; now apply Nat.lt_succ_r. Qed. 21 | 22 | Lemma le_plus_minus : forall n m : nat, n <= m -> m = n + (m - n). 23 | Proof. intros n m H; now rewrite Nat.add_comm, Nat.sub_add. Qed. 24 | 25 | Lemma le_plus_minus_r : forall n m : nat, n <= m -> n + (m - n) = m. 26 | Proof. intros n m H; now rewrite Nat.add_comm, Nat.sub_add. Qed. 27 | 28 | Lemma le_lt_or_eq : forall n m : nat, n <= m -> n < m \/ n = m. 29 | Proof. intros n m H; now apply Nat.lt_eq_cases. Qed. 30 | 31 | Lemma lt_n_Sm_le : forall n m : nat, n < S m -> n <= m. 32 | Proof. intros n m H; now apply Nat.lt_succ_r. Qed. 33 | 34 | Lemma lt_le_S : forall n m : nat, n < m -> S n <= m. 35 | Proof. intros n m H; now apply Nat.le_succ_l. Qed. 36 | 37 | Lemma lt_not_le : forall n m : nat, n < m -> ~ m <= n. 38 | Proof. intros n m H; now apply Nat.lt_nge. Qed. 39 | 40 | Lemma le_not_lt : forall n m : nat, n <= m -> ~ m < n. 41 | Proof. intros n m H; now apply Nat.le_ngt. Qed. 42 | 43 | (* Other compatibility lemmas (not needed after 8.17 *) 44 | 45 | (* NOTE: needed because Nat.le_add_l was only introduced in Coq 8.17 (see 46 | * Coq.Numbers.Natural.Abstract.NAddOrder). This can go away and be replaced 47 | * by Nat.le_add_r once the minimal Coq supported version is >= 8.17. *) 48 | Lemma le_add_l : forall n m : nat, n <= m + n. 49 | Proof. 50 | intros n m. rewrite <-(Nat.add_0_l n) at 1. apply Nat.add_le_mono. 51 | - exact (Nat.le_0_l m). 52 | - exact (Nat.le_refl n). 53 | Qed. 54 | 55 | (* NOTE: When the minimal supported Coq version is >= 8.16, 56 | remove it and rename NatUtil.Even_double into Nat.Even_double *) 57 | Lemma Even_double : forall n : nat, Nat.Even n -> n = Nat.double (Nat.div2 n). 58 | Proof. now intros n [k ->]; rewrite Nat.double_twice, Nat.div2_double. Qed. 59 | 60 | (* NOTE: When the minimal supported Coq version is >= 8.16, 61 | remove it and rename NatUtil.Odd_double into Nat.Odd_double *) 62 | Lemma Odd_double : forall n : nat, Nat.Odd n -> n = S (Nat.double (Nat.div2 n)). 63 | Proof. 64 | intros n [k ->]. 65 | now rewrite Nat.add_1_r, Nat.div2_succ_double, Nat.double_twice. 66 | Qed. 67 | 68 | (* NOTE: When the minimal supported Coq version is >= 8.16, 69 | remove it and rename NatUtil.Odd_double into Nat.Odd_double *) 70 | Lemma Even_div2 : forall n : nat, Nat.Even n -> Nat.div2 n = Nat.div2 (S n). 71 | Proof. now intros n [k ->]; rewrite Nat.div2_double, Nat.div2_succ_double. Qed. 72 | 73 | (* NOTE: When the minimal supported Coq version is >= 8.16, 74 | remove it and rename NatUtil.Odd_div2 into Nat.Odd_div2 *) 75 | Lemma Odd_div2 : forall n : nat, Nat.Odd n -> S (Nat.div2 n) = Nat.div2 (S n). 76 | Proof. 77 | intros n [k ->]; rewrite Nat.add_1_r, Nat.div2_succ_double. 78 | rewrite <-(Nat.add_1_r (S (2 * k))), (Nat.add_succ_comm (2 * k)). 79 | rewrite <-(Nat.mul_1_r 2) at 2. 80 | rewrite <-(Nat.mul_add_distr_l 2), Nat.add_succ_r. 81 | now rewrite Nat.add_0_r, Nat.div2_double. 82 | Qed. 83 | -------------------------------------------------------------------------------- /Coccinelle/examples/cime_trace/more_list_extention.v: -------------------------------------------------------------------------------- 1 | From CoLoR Require Import more_list decidable_set list_set. 2 | From Stdlib Require Import List Lia FunInd. 3 | From Stdlib Require Recdef. 4 | 5 | Module Type S. 6 | Parameter A : Type. 7 | End S. 8 | 9 | Module Make(AX:S)(X:decidable_set.ES with Definition A:=AX.A with Definition eq_A:=@eq AX.A). 10 | 11 | Module XSet := list_set.Make(X). 12 | 13 | Import XSet. 14 | Import X. 15 | Import AX. 16 | 17 | Lemma mem_bool_app : 18 | forall x l1 l2, 19 | (mem_bool eq_bool x (l1++l2)) = orb (mem_bool eq_bool x l1) (mem_bool eq_bool x l2). 20 | Proof. 21 | intros x l1. 22 | functional induction (mem_bool eq_bool x l1). 23 | 24 | simpl;reflexivity. 25 | simpl. 26 | intros l2;rewrite IHb0. auto with bool. 27 | Qed. 28 | 29 | Section S. 30 | Hypothesis order : A -> A -> bool. 31 | 32 | Function split (l : list A) (x: A) {struct l} : list A * list A := 33 | match l with 34 | | nil => (nil,nil) 35 | | y::l' => 36 | let (l1,l2) := split l' x in 37 | if order y x 38 | then (y::l1,l2) 39 | else (l1,y::l2) 40 | end. 41 | 42 | Lemma split_length : 43 | forall x l, 44 | let (l1,l2) := split l x in 45 | List.length l1 + List.length l2 = List.length l. 46 | Proof. 47 | clear. 48 | intros x l. 49 | functional induction (split l x). 50 | reflexivity. 51 | simpl. 52 | rewrite e0 in IHp. lia. 53 | simpl;rewrite e0 in IHp. lia. 54 | Qed. 55 | Import Recdef. 56 | Function qs (l:list A) {measure length } : list A := 57 | match l with 58 | | nil => nil 59 | | x::l' => 60 | let (l1,l2) := split l' x in 61 | (qs l1)++x::(qs l2) 62 | end. 63 | Proof. 64 | abstract(intros;generalize (split_length x l'); 65 | rewrite teq0; simpl; intros; lia). 66 | abstract(intros;generalize (split_length x l'); 67 | rewrite teq0; simpl; intros; lia). 68 | Defined. 69 | 70 | Lemma split_mem_bool : 71 | forall x y l l1 l2, split l x = (l1,l2) -> 72 | mem_bool eq_bool y l = orb (mem_bool eq_bool y l1) (mem_bool eq_bool y l2). 73 | Proof. 74 | intros x y l; 75 | functional induction (split l x); 76 | intros l1' l2' H;injection H;clear H;intros h1 h2;subst;simpl. 77 | reflexivity. 78 | rewrite IHp with (1:=e0);auto with bool. 79 | rewrite IHp with (1:=e0);auto with bool. 80 | case (eq_bool y y0);case (mem_bool eq_bool y l1'); case (mem_bool eq_bool y l2); 81 | reflexivity. 82 | Qed. 83 | 84 | Lemma qs_valid : forall f l, 85 | more_list.mem_bool eq_bool f l = more_list.mem_bool eq_bool f (qs l). 86 | Proof. 87 | intros f l. 88 | functional induction (qs l). 89 | reflexivity. 90 | simpl. 91 | rewrite mem_bool_app. 92 | simpl. 93 | rewrite <- IHl0. 94 | rewrite <- IHl1. 95 | rewrite split_mem_bool with (1:=e0). 96 | case (eq_bool f x);case (mem_bool eq_bool f l1); case (mem_bool eq_bool f l2); 97 | reflexivity. 98 | Qed. 99 | End S. 100 | 101 | Lemma change_in : 102 | forall order f l, 103 | In f l -> 104 | more_list.mem_bool eq_bool f (qs order (remove_red l)) = true. 105 | Proof. 106 | intros order f l H. 107 | apply (more_list.in_impl_mem (@eq A) 108 | (fun a => eq_refl a)) in H. 109 | apply XSet.remove_red_included in H. 110 | generalize (more_list.mem_bool_ok _ _ eq_bool_ok f (remove_red l)). 111 | case_eq (more_list.mem_bool eq_bool f (remove_red l)). 112 | clear H;intros H _. 113 | 2:clear -H;intros _ abs;contradiction. 114 | rewrite (qs_valid order) in H. 115 | exact H. 116 | Qed. 117 | End Make. 118 | -------------------------------------------------------------------------------- /Util/Relation/Preorder.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Solange Coupet-Grimal and William Delobel, 2006-01-09 6 | 7 | Various properties of preorders 8 | *) 9 | 10 | From CoLoR Require Import LogicUtil. 11 | From Stdlib Require Import Relations. 12 | 13 | Section PreOrderFacts. 14 | Variable A : Type. 15 | Variable leA : relation A. 16 | 17 | Definition eqA : relation A := fun (f g : A) => leA f g /\ leA g f. 18 | Definition ltA : (relation A) := fun (f g : A) => leA f g /\ ~ eqA f g. 19 | 20 | Variable leA_preorder : preorder A leA. 21 | 22 | Lemma eqA_equivalence : equivalence A eqA. 23 | Proof. 24 | inversion leA_preorder. 25 | split. 26 | (* reflexive : *) 27 | intro a; split; apply preord_refl. 28 | (* transitive : *) 29 | intros a b c H1 H2. 30 | elim H1; clear H1; intros a_le_g b_le_a. 31 | elim H2; clear H2; intros b_le_c c_le_b. 32 | split. 33 | apply preord_trans with b; hyp. 34 | apply preord_trans with b; hyp. 35 | (* symmetric : *) 36 | intros a b H. 37 | elim H; split; hyp. 38 | Qed. 39 | 40 | Lemma ltA_antisym : forall (a b : A), ltA a b -> ltA b a -> False. 41 | Proof. 42 | intros a b H1 H2. 43 | elim H1; clear H1; intros a_le_b b_neq_a. 44 | elim H2; clear H2; intros b_le_a a_neq_b. 45 | apply a_neq_b; split; hyp. 46 | Qed. 47 | 48 | Lemma ltA_trans : transitive A ltA. 49 | Proof. 50 | inversion leA_preorder. 51 | intros a b c H1 H2. 52 | elim H1; clear H1; intros a_le_b a_neq_b. 53 | elim H2; clear H2; intros b_le_c b_neq_c. 54 | split. 55 | apply preord_trans with b; hyp. 56 | intro H. 57 | elim H; clear H; intros a_le_c c_le_a. 58 | apply a_neq_b; split. 59 | hyp. 60 | apply preord_trans with c; hyp. 61 | Qed. 62 | 63 | Lemma leA_dec_to_eqA_dec : (forall (a b : A), leA a b \/ ~ leA a b) -> 64 | forall (a b : A), eqA a b \/ ~ eqA a b. 65 | Proof. 66 | intros leA_dec a b; elim (leA_dec a b); intro case_a_le_b. 67 | elim (leA_dec b a); intro case_b_le_a; [left | right]. 68 | split; trivial. 69 | intro H; elim H; clear H; intros H1 H2. 70 | apply case_b_le_a; trivial. 71 | right; intro H; apply case_a_le_b. 72 | elim H; trivial. 73 | Qed. 74 | 75 | Lemma ltA_eqA_compat_r : forall (a b b' :A), eqA b b' -> ltA a b -> ltA a b'. 76 | Proof. 77 | inversion leA_preorder. 78 | intros a b b' b_eq a_lt_b. 79 | elim a_lt_b; clear a_lt_b; intros a_le_b a_neq_b. 80 | split. 81 | elim b_eq; clear b_eq; intros b_le_b' b'_le_b. 82 | apply preord_trans with b; hyp. 83 | elim (eqA_equivalence); intros eqA_refl eqA_trans eqA_sym. 84 | intro; apply a_neq_b; apply eqA_trans with b'; trivial. 85 | elim b_eq; split; trivial. 86 | Qed. 87 | 88 | Lemma ltA_eqA_compat_l : forall (a b b' :A), eqA b b' -> ltA b a -> ltA b' a. 89 | Proof. 90 | inversion leA_preorder. 91 | intros a b b' b_eq b_lt_a. 92 | elim b_lt_a; clear b_lt_a; intros b_le_a b_neq_a. 93 | split. 94 | elim b_eq; clear b_eq; intros b_le_b' b'_le_b. 95 | apply preord_trans with b; hyp. 96 | elim (eqA_equivalence); intros eqA_refl eqA_trans eqA_sym. 97 | intro; apply b_neq_a; apply eqA_trans with b'; trivial. 98 | Qed. 99 | 100 | Section Decidability. 101 | 102 | Variable leA_dec : forall a b, {leA a b} + {~leA a b}. 103 | 104 | Lemma eqA_dec : forall a b, {eqA a b} + {~eqA a b}. 105 | 106 | Proof. 107 | intros. unfold eqA. 108 | destruct (leA_dec a b); intuition. 109 | destruct (leA_dec b a); intuition. 110 | Defined. 111 | 112 | Lemma ltA_dec : forall a b, {ltA a b} + {~ltA a b}. 113 | 114 | Proof. 115 | intros. unfold ltA. 116 | destruct (leA_dec a b); intuition. 117 | destruct (eqA_dec a b); intuition. 118 | Defined. 119 | 120 | End Decidability. 121 | 122 | End PreOrderFacts. 123 | -------------------------------------------------------------------------------- /Term/SimpleType/TermsBuilding.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2006-04-27 6 | 7 | Constructing terms. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import RelExtras ListExtras LogicUtil. 13 | From CoLoR Require TermsActiveEnv. 14 | From Stdlib Require Import Lia. 15 | 16 | Module TermsBuilding (Sig : TermsSig.Signature). 17 | 18 | Module Export TAE := TermsActiveEnv.TermsActiveEnv Sig. 19 | 20 | Record appCond : Type := { 21 | appL: Term; 22 | appR: Term; 23 | eqEnv: env appL = env appR; 24 | typArr: isArrowType (type appL); 25 | typOk: type_left (type appL) = type appR 26 | }. 27 | 28 | Definition buildApp : appCond -> Term. 29 | 30 | Proof. 31 | intro t; inversion t as [L R eq_env typ_arr typ_ok]. 32 | destruct L as [?? typeL typingL]; destruct R as [??? typingR]; simpl in *. 33 | rewrite eq_env in typingL. 34 | destruct typeL; try contr; simpl in *. 35 | rewrite typ_ok in typingL. 36 | exact (buildT (TApp typingL typingR)). 37 | Defined. 38 | 39 | Lemma buildApp_isApp : forall a, isApp (buildApp a). 40 | 41 | Proof. 42 | intros; destruct a; term_type_inv appL0; term_type_inv appR0. 43 | Qed. 44 | 45 | Lemma buildApp_Lok : forall a, appBodyL (buildApp_isApp a) = a.(appL). 46 | 47 | Proof. 48 | destruct a; destruct appR0; term_type_inv appL0. 49 | Qed. 50 | 51 | Lemma buildApp_Rok : forall a, appBodyR (buildApp_isApp a) = a.(appR). 52 | 53 | Proof. 54 | destruct a; destruct appR0; term_type_inv appL0. 55 | Qed. 56 | 57 | Lemma buildApp_preterm : forall a, 58 | term (buildApp a) = term a.(appL) @@ term a.(appR). 59 | 60 | Proof. 61 | destruct a; destruct appR0; term_type_inv appL0. 62 | Qed. 63 | 64 | Lemma buildApp_env_l : forall a, env (buildApp a) = env a.(appL). 65 | 66 | Proof. 67 | destruct a; destruct appR0; term_type_inv appL0. 68 | Qed. 69 | 70 | Lemma buildApp_type : forall a, 71 | type (buildApp a) = type_right (type a.(appL)). 72 | 73 | Proof. 74 | destruct a; destruct appR0; term_type_inv appL0. 75 | Qed. 76 | 77 | Record absCond : Type := { 78 | absB: Term; 79 | absT: SimpleType; 80 | envNotEmpty: env absB |= 0 := absT 81 | }. 82 | 83 | Definition buildAbs : absCond -> Term. 84 | 85 | Proof. 86 | intro t; inversion t as [aBody aType envCond]. 87 | destruct aBody as [env ?? typing]; simpl in *; destruct env. 88 | try_solve. 89 | destruct o. 90 | exact (buildT (TAbs typing)). 91 | try_solve. 92 | Defined. 93 | 94 | Lemma buildAbs_isAbs: forall a, isAbs (buildAbs a). 95 | Proof. 96 | destruct a as [[env ???] ??]; destruct env. 97 | try_solve. 98 | destruct o; try_solve. 99 | Qed. 100 | 101 | Lemma buildAbs_absBody : forall a, absBody (buildAbs_isAbs a) = a.(absB). 102 | 103 | Proof. 104 | destruct a as [[env ???] ??]; destruct env. 105 | try_solve. 106 | destruct o; try_solve. 107 | Qed. 108 | 109 | Lemma buildAbs_absType : forall a, absType (buildAbs_isAbs a) = a.(absT). 110 | 111 | Proof. 112 | destruct a as [[env ???] ??]; destruct env. 113 | try_solve. 114 | destruct o; try_solve. 115 | unfold VarD in * . 116 | inversion envNotEmpty0; trivial. 117 | Qed. 118 | 119 | Lemma buildAbs_env : forall a, env (buildAbs a) = tail (env a.(absB)). 120 | 121 | Proof. 122 | destruct a as [[env ???] ??]; destruct env. 123 | try_solve. 124 | destruct o; try_solve. 125 | Qed. 126 | 127 | Definition buildVar : forall A x, (copy x None ++ A [#] EmptyEnv) |- %x := A. 128 | 129 | Proof. 130 | constructor; unfold VarD. 131 | rewrite nth_app_right; autorewrite with datatypes using try lia. 132 | replace (x - x) with 0; trivial. 133 | lia. 134 | Defined. 135 | 136 | Lemma buildVar_minimal : forall A x, envMinimal (buildT (buildVar A x)). 137 | 138 | Proof. 139 | intros; unfold envMinimal; trivial. 140 | Qed. 141 | 142 | End TermsBuilding. 143 | -------------------------------------------------------------------------------- /Util/Vector/VecFilterPerm.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2009-09-09 6 | 7 | vector filtering with permutations 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import VecUtil ListUtil NatUtil LogicUtil ListNodup BoundNat. 13 | 14 | (***********************************************************************) 15 | (** filtering function *) 16 | 17 | Fixpoint Vfilter A n (l : list (N n)) (v : vector A n) : vector A (length l) := 18 | match l as l return vector A (length l) with 19 | | nil => Vnil 20 | | x :: l' => Vcons (Vnth v x) (Vfilter l' v) 21 | end. 22 | 23 | (***********************************************************************) 24 | (** properties *) 25 | 26 | Lemma Vfilter_eq_notin : forall A n (l : list (N n)) i (vi : vector A i) t u 27 | j (vj : vector A j) (e : i+S j=n), (forall hi : i 28 | Vfilter l (Vcast (Vapp vi (Vcons t vj)) e) 29 | = Vfilter l (Vcast (Vapp vi (Vcons u vj)) e). 30 | 31 | Proof. 32 | induction l; simpl; intros. refl. apply Vcons_eq. split. 33 | rewrite !Vnth_cast, !Vnth_app. 34 | destruct a as [k h]. unfold N_val; simpl proj1_sig. case (le_gt_dec i k); intro. 35 | rewrite !Vnth_cons. case (lt_ge_dec 0 (k-i)); intro. refl. 36 | assert (k=i). lia. subst k. ded (H h). fo. 37 | refl. apply IHl. fo. 38 | Qed. 39 | 40 | Lemma Vfilter_app : forall A n (l1 l2 : list (N n)) (v : vector A n), 41 | Vfilter (l1 ++ l2) v = Vcast (Vapp (Vfilter l1 v) (Vfilter l2 v)) 42 | (sym_eq (length_app l1 l2)). 43 | 44 | Proof. 45 | induction l1; simpl; intros. rewrite Vcast_refl. refl. 46 | rewrite Vcast_cons. f_equal. rewrite (IHl1 l2 v). apply Vcast_pi. 47 | Qed. 48 | 49 | Lemma Vfilter_app_eq : forall A n (l l1 l2 : list (N n)) (v : vector A n) 50 | (e : length l1+length l2 = length l), 51 | l=l1++l2 -> Vfilter l v = Vcast (Vapp (Vfilter l1 v) (Vfilter l2 v)) e. 52 | 53 | Proof. intros. subst l. rewrite Vfilter_app. apply Vcast_pi. Qed. 54 | 55 | Arguments Vfilter_app_eq [A n l l1 l2 v] _ _. 56 | 57 | Lemma Vfilter_eq_in : forall A n (v : vector A n) (l : list (N n)) 58 | (rf : nodup (map (@N_val n) l)) i, In i (map (@N_val n) l) -> 59 | exists hi : i B) n (v : vector A n) (l : list (N n)), 77 | Vfilter l (Vmap f v) = Vmap f (Vfilter l v). 78 | 79 | Proof. induction l; simpl; intros. refl. rewrite IHl, Vnth_map. refl. Qed. 80 | 81 | Lemma Vin_filter_elim_in : forall A n (l : list (N n)) (v : vector A n) x, 82 | Vin x (Vfilter l v) -> Vin x v. 83 | 84 | Proof. 85 | induction l; simpl; intros. contr. destruct H. subst. apply Vnth_in. 86 | apply IHl. hyp. 87 | Qed. 88 | 89 | Lemma Vin_filter_elim_nth : forall A n (l : list (N n)) (v : vector A n) x, 90 | Vin x (Vfilter l v) -> exists i (hi : i Vin (Vnth v hi) (Vfilter l v). 100 | 101 | Proof. 102 | induction l; simpl; intros. hyp. destruct a. destruct H. 103 | left. apply Vnth_eq. inversion H. subst. refl. 104 | right. apply IHl. hyp. 105 | Qed. 106 | -------------------------------------------------------------------------------- /Term/WithArity/AWFMInterpretation.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2006-10-19 6 | - Sebastien Hinderer, 2004-02-25 7 | 8 | well-founded monotone interpretations 9 | *) 10 | 11 | Set Implicit Arguments. 12 | 13 | From CoLoR Require Export AInterpretation. 14 | From CoLoR Require Import ATerm RelUtil ASubstitution NaryFunction AContext 15 | VecUtil SN ARelation LogicUtil. 16 | From Stdlib Require Import Lia PeanoNat. 17 | 18 | Section S. 19 | 20 | Variable Sig : Signature. 21 | 22 | Notation term := (term Sig). Notation terms := (vector term). 23 | 24 | Variable I : interpretation Sig. 25 | 26 | Section IR. 27 | 28 | Variable R : relation I. 29 | 30 | Definition IR : relation term := 31 | fun t u => forall xint, R (term_int xint t) (term_int xint u). 32 | 33 | (***********************************************************************) 34 | (** properties of IR *) 35 | 36 | Lemma IR_reflexive : reflexive R -> reflexive IR. 37 | 38 | Proof. 39 | intro. unfold reflexive, IR. auto. 40 | Qed. 41 | 42 | Lemma IR_transitive : transitive R -> transitive IR. 43 | 44 | Proof. 45 | intro. unfold transitive, IR. intros. exact (H _ _ _ (H0 xint) (H1 xint)). 46 | Qed. 47 | 48 | Lemma IR_substitution_closed : substitution_closed IR. 49 | 50 | Proof. 51 | unfold transp, substitution_closed, IR. intros t1 t2 s H xint0. 52 | rewrite !substitution_lemma. apply (H (beta xint0 s)). 53 | Qed. 54 | 55 | Definition monotone := forall f, Vmonotone1 (fint I f) R. 56 | 57 | Lemma IR_context_closed : monotone -> context_closed IR. 58 | 59 | Proof. 60 | unfold transp, context_closed, IR. intros. 61 | gen (H0 xint). clear H0. intro. induction c. 62 | simpl. exact H0. 63 | simpl fill. simpl. 64 | do 2 (rewrite Vmap_cast, Vmap_app). simpl. apply H. exact IHc. 65 | Qed. 66 | 67 | Lemma IR_WF : WF R -> WF IR. 68 | 69 | Proof. 70 | intro. set (xint := fun x:nat => some_elt I). 71 | apply (WF_incl (fun t1 t2 => R (term_int xint t1) (term_int xint t2))). 72 | unfold Basics.flip, inclusion. auto. apply (WF_inverse (term_int xint) H). 73 | Qed. 74 | 75 | Lemma IR_reduction_ordering : monotone -> WF R -> reduction_ordering IR. 76 | 77 | Proof. 78 | split. apply IR_WF. exact H0. split. apply IR_substitution_closed. 79 | apply IR_context_closed. exact H. 80 | Qed. 81 | 82 | (***********************************************************************) 83 | (** equivalent definition *) 84 | 85 | Definition IR' : relation term := fun t u => 86 | let n := 1 + max (maxvar t) (maxvar u) in 87 | forall v : vector I n, 88 | let xint := val_of_vec I v in 89 | R (term_int xint t) (term_int xint u). 90 | 91 | Lemma IR_incl_IR' : IR << IR'. 92 | 93 | Proof. 94 | unfold inclusion, IR, IR'. intros. apply H. 95 | Qed. 96 | 97 | Lemma IR'_incl_IR : IR' << IR. 98 | 99 | Proof. 100 | unfold inclusion, IR, IR'. intros. set (m := max (maxvar x) (maxvar y)). 101 | assert (maxvar x <= m). unfold m. apply Nat.le_max_l. 102 | assert (maxvar y <= m). unfold m. apply Nat.le_max_r. 103 | assert (term_int xint x = term_int (fval xint (1+m)) x). 104 | apply term_int_eq_fval_lt. lia. rewrite H2. 105 | assert (term_int xint y = term_int (fval xint (1+m)) y). 106 | apply term_int_eq_fval_lt. lia. rewrite H3. 107 | unfold fval. apply H. 108 | Qed. 109 | 110 | Lemma IR_eq_IR' : IR == IR'. 111 | 112 | Proof. 113 | split. exact IR_incl_IR'. exact IR'_incl_IR. 114 | Qed. 115 | 116 | End IR. 117 | 118 | (***********************************************************************) 119 | (** monotony wrt R *) 120 | 121 | Section inclusion. 122 | 123 | Variables R R' : relation I. 124 | 125 | Lemma IR_inclusion : R << R' -> IR R << IR R'. 126 | 127 | Proof. 128 | intro. unfold inclusion, IR. intros. apply H. apply H0. 129 | Qed. 130 | 131 | End inclusion. 132 | 133 | End S. 134 | -------------------------------------------------------------------------------- /Term/WithArity/ASubterm.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Sidi Ould-Biha, 2010-04-27 6 | 7 | Properties of the subterm relation. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import AContext VecUtil ListUtil LogicUtil NatUtil RelUtil 13 | ASN. 14 | 15 | Section S. 16 | 17 | Variable Sig : Signature. 18 | 19 | Notation term := (term Sig). Notation terms := (vector term). 20 | 21 | (***********************************************************************) 22 | (** list of immediate subterms *) 23 | 24 | Fixpoint subterm_lst t := 25 | match t with 26 | | Var _ => nil 27 | | Fun f ts => 28 | let fix subterm_lst_vec k (us : terms k) {struct us} : list term := 29 | match us with 30 | | Vnil => nil 31 | | Vcons u1 us' => subterm_lst u1 ++ subterm_lst_vec _ us' 32 | end 33 | in list_of_vec ts ++ subterm_lst_vec (arity f) ts 34 | end. 35 | 36 | Fixpoint subterm_lst_vec k (us : terms k) : list term := 37 | match us with 38 | | Vnil => nil 39 | | Vcons u1 us' => subterm_lst u1 ++ subterm_lst_vec us' 40 | end. 41 | 42 | Lemma subterm_lst_fun : forall F ts, 43 | subterm_lst (Fun F ts) = list_of_vec ts ++ subterm_lst_vec ts. 44 | 45 | Proof. 46 | refl. 47 | Qed. 48 | 49 | (***********************************************************************) 50 | (** list of immediate subterms *) 51 | 52 | Lemma In_subterm_lst_vec_elim : forall v n (ts : terms n), 53 | In v (subterm_lst_vec ts) -> 54 | exists i, exists p : i < n, In v (subterm_lst (Vnth ts p)). 55 | 56 | Proof. 57 | induction ts. contr. 58 | intros In_v; simpl in In_v. rewrite in_app in In_v. intuition. 59 | exists 0. exists (Nat.lt_0_succ n). simpl. hyp. 60 | destruct H0 as [i H0]. destruct H0 as [p H0]. 61 | assert (p' : S i < S n). lia. exists (S i). exists p'. simpl. 62 | assert (H1 : (NatCompat.lt_S_n p') = p). apply lt_unique. rewrite H1. hyp. 63 | Qed. 64 | 65 | Lemma In_subterm_lst_vec_intro : forall n (ts : terms n) v i (p : i < n), 66 | In v (subterm_lst (Vnth ts p)) -> In v (subterm_lst_vec ts). 67 | 68 | Proof. 69 | induction ts. intros. lia. 70 | destruct i; simpl; intros; rewrite in_app. 71 | left; hyp. right. apply (IHts _ _ (NatCompat.lt_S_n p)); auto. 72 | Qed. 73 | 74 | (***********************************************************************) 75 | (** the supterm relation is finitely branching *) 76 | 77 | Notation supterm := (@supterm Sig). 78 | 79 | Lemma fin_branch_supterm : finitely_branching supterm. 80 | 81 | Proof. 82 | intros x. exists (subterm_lst x). pattern x; apply term_ind_forall; clear x. 83 | simpl. intros v y. split; try tauto. intro HS; destruct HS as [c Hc]. 84 | destruct (var_eq_fill (proj2 Hc)) as [Hc' _]; intuition. 85 | intros f ts Hts y. split; rewrite subterm_lst_fun, in_app. 86 | 2:{ intuition. gen (in_list_of_vec H0). apply subterm_fun. 87 | ded (In_subterm_lst_vec_elim _ _ H0). decomp H. 88 | ded (Vforall_nth x0 Hts y). apply (@subterm_trans _ _ (Vnth ts x0)). 89 | apply H. hyp. apply subterm_fun. apply Vnth_in. } 90 | intros. destruct H as [c Hc]. destruct (fun_eq_fill (proj2 Hc)). intuition. 91 | destruct H as [i H]. destruct H as [j H]. destruct H as [r H]. 92 | destruct H as [ti H]. destruct H as [c0 H]. destruct H as [tj Hc0]. 93 | destruct Hc as [Hc Hf]. rewrite Hc0 in Hf. simpl in Hf. Funeqtac. 94 | destruct c0. simpl in H. left. rewrite H. apply list_of_vec_in. 95 | apply Vin_cast_intro. apply Vin_appr. simpl. left; refl. 96 | right. assert (p : i < arity f). rewrite <- r. lia. 97 | assert (Hs : supterm (Vnth ts p) y). 98 | rewrite H, Vnth_cast, Vnth_app. 99 | destruct (le_gt_dec i i). 2: lia. 100 | set (q := (Vnth_app_aux (S j) (Vnth_cast_aux r p) l)). 101 | rewrite (Vnth_eq _ q (Nat.lt_0_succ j)); try lia. simpl. 102 | exists (Cont f0 e t c0 t0). simpl. intuition. discr. 103 | ded (Vforall_nth p Hts y). apply (In_subterm_lst_vec_intro _ _ p). 104 | apply (proj1 H0). auto. 105 | Qed. 106 | 107 | End S. 108 | -------------------------------------------------------------------------------- /CONTRIBUTING: -------------------------------------------------------------------------------- 1 | Guideline for CoLoR developers 2 | ------------------------------ 3 | 4 | - For new contributors, update the COPYRIGHT file. 5 | 6 | - Don't forget to update the CHANGES and THANKS files. 7 | 8 | - If you bring important modifications to an already existing file, 9 | add in the header your name with the modification date yy-mm-dd. 10 | 11 | How to create a new release: 12 | ---------------------------- 13 | 14 | In ~/color-svn: 15 | - check that every compiles: make clean-all && make && make doc 16 | - update CHANGES, README and INSTALL 17 | - make install-doc 18 | - make install-dist 19 | - commit your changes 20 | - ./create_dist $new_version_name (e.g. 1.4.0) 21 | (it computes md5sum color.$new_version_name) 22 | - check that color.$new_version_name compiles 23 | - add color.$new_version_name on https://gforge.inria.fr/frs/?group_id=467 24 | 25 | In ~/src/opam-coq-archive-fblanqui-git: 26 | - update-from-remote 27 | - cd released/packages/coq-color 28 | - cp coq-color.$old_version_name coq-color.$new_version_name 29 | - cd coq-color.$new_version_name 30 | - update descr, opam and url 31 | - nettoie 32 | - cd .. 33 | - add coq-color.$new_version_name 34 | - ci -m "$comment" 35 | 36 | On https://github.com/fblanqui/opam-coq-archive: 37 | - create a pull request 38 | 39 | In ~/rewriting-svn/web: 40 | - update color/index.html 41 | - make 42 | - make install 43 | 44 | After pull request approval, send a mail to color@inria.fr. 45 | 46 | About compilation: 47 | ------------------ 48 | 49 | - Before any commit, make sure that CoLoR compiles with the last 50 | stable release of Coq and, if possible, with the development version 51 | of Coq. Check also that Rainbow examples still work. 52 | 53 | - Put useful messages when you commit. 54 | 55 | - When you add a file, do "make config" to build Makefile.Coq and the 56 | dependencies again. 57 | 58 | About the CoLoR style: 59 | ---------------------- 60 | 61 | - Try to follow the style currently used in CoLoR. 62 | 63 | - Use maximum 78 characters per line. 64 | 65 | - Add the CoLoR header at the beginning of each file with the creation 66 | date yy-mm-dd. 67 | 68 | - Only use Parameter, Definition, Lemma (no Axiom, Theorem, etc.). 69 | 70 | - Put comments! And use the possibilities offered by coqdoc. For 71 | instance, "(** " creates a section title. Check the result by doing 72 | "make doc". 73 | 74 | - Try to use the Coq standard library as much as possible except for 75 | the definition of well-foundedness: it is redefined in 76 | Util/Relation/SN.v. 77 | 78 | - Definitions and lemmas about structures of general interest must be 79 | put in Util. Term structures are in Term and every termination 80 | technique must be put in a new directory at the root of the project. 81 | 82 | - Files about algebraic terms (Term/WithArity) should be prefixed with 83 | A. Files about varyadic terms (Term/Varyadic) should be prefixed with 84 | V. 85 | 86 | - Define decidability lemmas using boolean function. 87 | 88 | - Do not use setoid rewrites in decidability lemmas. 89 | 90 | - Try to build the names of lemmas by concatening its main 91 | ingredients. For instance: 92 | 93 | Lemma Vforall_in : forall P x n (v : vec n), Vforall P v -> Vin x v -> P x. 94 | 95 | You can also add the suffix "intro" or "elim" for specifying its 96 | use. For instance: 97 | 98 | Lemma Vforall_intro : forall (P : A->Prop) n (v : vec n), 99 | (forall x, Vin x v -> P x) -> Vforall P v. 100 | 101 | provides a way of proving Vforall (introduction rule). 102 | 103 | Lemma Vin_elim : forall x n (v : vec n), 104 | Vin x v -> exists n1, exists v1 : vec n1, exists n2, exists v2 : vec n2, 105 | exists H : n1 + S n2 = n, v = Vcast (Vapp v1 (Vcons x v2)) H. 106 | 107 | provides a way to deduce some information from Vin (elimination rule). 108 | 109 | - Use "Require Export" as less as possible. Use "Require Import" 110 | instead. This reduces the compilation time. 111 | 112 | - Use (*COQ:...*) for indicating problems with Coq, 113 | (*FIXME:...*) for indicating something to be fixed, 114 | (*REMARK:...*) for adding some important remark 115 | (e.g. something that could be improved), 116 | (*REMOVE:...*) for indicating something to be removed, 117 | (*MOVE:...*) for indicating something to be moved. 118 | -------------------------------------------------------------------------------- /HORPO/HorpoComp.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski, 2006-04-27 6 | 7 | Some computability results instantiated for horpo. 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import RelExtras ListExtras Computability Horpo. 13 | From Stdlib Require Import Setoid Morphisms. 14 | 15 | Module HorpoComp (S : TermsSig.Signature) 16 | (Prec : Horpo.Precedence with Module S := S). 17 | 18 | Module Export Comp := Computability.Computability S Prec. 19 | 20 | Import List. 21 | 22 | Definition horpo_lt := transp horpo. 23 | Definition horpo_mul_lt := MultisetLT horpo. 24 | 25 | Notation "X << Y" := (horpo_lt X Y) (at level 55). 26 | Definition Htrans := clos_trans horpo. 27 | 28 | Definition Terms := list Term. 29 | Definition CompH := Computable horpo. 30 | Definition CompTerms (Ts: Terms) := AllComputable horpo Ts. 31 | Definition CompSubst (G: Subst) := forall Q, In (Some Q) G -> CompH Q. 32 | 33 | Definition WFterms := { Ts: Terms | CompTerms Ts }. 34 | Definition WFterms_to_mul (Ts: WFterms) := list2multiset (proj1_sig Ts). 35 | Coercion WFterms_to_mul: WFterms >-> Multiset. 36 | Definition H_WFterms_lt (M N: WFterms) := horpo_mul_lt M N. 37 | 38 | #[global] Hint Unfold horpo_lt horpo_mul_lt CompH CompTerms : horpo. 39 | 40 | Lemma horpo_comp_imp_acc M : CompH M -> AccR horpo M. 41 | 42 | Proof. intro Mcomp. apply comp_imp_acc; eauto with horpo. Qed. 43 | 44 | Lemma horpo_comp_step_comp M N : CompH M -> M >> N -> CompH N. 45 | 46 | Proof. 47 | intros Mcomp MN. 48 | unfold CompH; apply comp_step_comp with M; eauto with horpo. 49 | Qed. 50 | 51 | Lemma horpo_comp_manysteps_comp M N : CompH M -> M >>* N -> CompH N. 52 | 53 | Proof. 54 | intros Mcomp M_N. 55 | unfold CompH; apply comp_manysteps_comp with M; eauto with horpo. 56 | Qed. 57 | 58 | Lemma horpo_comp_pflat N Ns : isPartialFlattening Ns N -> algebraic N -> 59 | AllComputable horpo Ns -> CompH N. 60 | 61 | Proof. 62 | intros NsN Nnorm NsC; unfold CompH. apply comp_pflat with Ns; trivial. 63 | Qed. 64 | 65 | Lemma horpo_neutral_comp_step : forall M, algebraic M -> isNeutral M -> 66 | (CompH M <-> (forall N, M >> N -> CompH N)). 67 | 68 | Proof. 69 | intros M Mneutral; unfold CompH. 70 | apply neutral_comp_step; eauto with horpo. 71 | Qed. 72 | 73 | Lemma CompH_morph_aux : forall x1 x2 : Term, x1 ~ x2 -> CompH x1 -> CompH x2. 74 | 75 | Proof. 76 | intros t t' teqt' H_t. 77 | unfold CompH. 78 | apply Computable_morph_aux with t; eauto with horpo. 79 | Qed. 80 | 81 | Global Instance CompH_morph : Proper (terms_conv ==> iff) CompH. 82 | 83 | Proof. 84 | intros; split; apply CompH_morph_aux; auto using terms_conv_sym. 85 | Qed. 86 | 87 | Lemma horpo_comp_conv: forall M M', CompH M -> M ~ M' -> CompH M'. 88 | 89 | Proof. intros. rewrite <- H0; trivial. Qed. 90 | 91 | Lemma horpo_var_comp : forall M, isVar M -> CompH M. 92 | 93 | Proof. 94 | intros M Mvar; unfold CompH. apply var_comp; eauto with horpo. 95 | Qed. 96 | 97 | Lemma horpo_comp_abs : forall M (Mabs: isAbs M), algebraic M -> 98 | (forall G (cs: correct_subst (absBody Mabs) G) T, 99 | isSingletonSubst T G -> CompH T -> 100 | CompH (subst cs)) -> CompH M. 101 | 102 | Proof. 103 | intros M Mnorm Mabs H; unfold CompH. eapply comp_abs; eauto with horpo. 104 | Qed. 105 | 106 | Lemma horpo_comp_lift : forall N, CompH N -> CompH (lift N 1). 107 | 108 | Proof. 109 | intros. 110 | setoid_replace (lift N 1) with N using relation terms_conv; trivial. 111 | apply terms_conv_sym; apply terms_lift_conv. 112 | Qed. 113 | 114 | Lemma horpo_comp_app : forall M (Mapp: isApp M), 115 | CompH (appBodyL Mapp) -> CompH (appBodyR Mapp) -> CompH M. 116 | 117 | Proof. 118 | intros M Mapp Ml Mr; unfold CompH. apply comp_app with Mapp; trivial. 119 | Qed. 120 | 121 | Lemma horpo_comp_algebraic : forall M, CompH M -> algebraic M. 122 | 123 | Proof. intros. apply comp_algebraic with horpo; trivial. Qed. 124 | 125 | Lemma horpo_comp_units_comp M : 126 | (forall N, isAppUnit N M -> CompH N) -> CompH M. 127 | 128 | Proof. intro Munits; unfold CompH. apply comp_units_comp; trivial. Qed. 129 | 130 | End HorpoComp. 131 | -------------------------------------------------------------------------------- /PolyInt/APolyInt_MA.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Adam Koprowski and Hans Zantema, 2007-04-25 6 | 7 | Polynomial interpretations in the setting of monotone algebras. 8 | *) 9 | 10 | From CoLoR Require Import APolyInt AMonAlg ZUtil RelUtil PositivePolynom ATrs 11 | MonotonePolynom LogicUtil BoolUtil. 12 | From Stdlib Require List. 13 | 14 | Module Type TPolyInt. 15 | Parameter sig : Signature. 16 | Parameter trsInt : PolyInterpretation sig. 17 | Parameter trsInt_wm : PolyWeakMonotone trsInt. 18 | End TPolyInt. 19 | 20 | Module PolyInt (Export PI : TPolyInt). 21 | 22 | (*FIXME: remove this layer*) 23 | Module Export MonotoneAlgebra <: MonotoneAlgebraType. 24 | 25 | Definition Sig := sig. 26 | 27 | Definition I := Int_of_PI trsInt_wm. 28 | 29 | Definition succ := Dgt. 30 | Definition succeq := Dge. 31 | 32 | Definition refl_succeq := refl_Dge. 33 | Definition trans_succ := trans_Dgt. 34 | Definition trans_succeq := trans_Dge. 35 | 36 | Lemma monotone_succeq : AWFMInterpretation.monotone I succeq. 37 | 38 | Proof. 39 | unfold I. apply pi_monotone_eq. 40 | Qed. 41 | 42 | Definition succ_wf := WF_Dgt. 43 | 44 | Lemma succ_succeq_compat : absorbs_left succ succeq. 45 | 46 | Proof. 47 | unfold succ, succeq. intros p q pq. destruct pq as [r [pr rq]]. 48 | unfold Dgt, Dlt, transp. apply Z.lt_le_trans with (val r); auto. 49 | Qed. 50 | 51 | Definition rulePoly_gt l r := rulePoly_gt trsInt (@mkRule sig l r). 52 | 53 | Definition succ' l r := coef_pos (rulePoly_gt l r). 54 | 55 | Definition bsucc' l r := bcoef_pos (rulePoly_gt l r). 56 | 57 | Lemma bsucc'_ok : forall l r, bsucc' l r = true <-> succ' l r. 58 | 59 | Proof. 60 | intros l r. unfold bsucc', succ'. apply bcoef_pos_ok. 61 | Qed. 62 | 63 | Lemma succ'_dec : rel_dec succ'. 64 | 65 | Proof. 66 | intros l r. unfold succ'. eapply dec. apply bcoef_pos_ok. 67 | Defined. 68 | 69 | Definition rulePoly_ge l r := rulePoly_ge trsInt (@mkRule sig l r). 70 | 71 | Definition succeq' l r := coef_pos (rulePoly_ge l r). 72 | 73 | Definition bsucceq' l r := bcoef_pos (rulePoly_ge l r). 74 | 75 | Lemma bsucceq'_ok : forall l r, bsucceq' l r = true <-> succeq' l r. 76 | 77 | Proof. 78 | intros l r. unfold bsucceq', succeq'. apply bcoef_pos_ok. 79 | Qed. 80 | 81 | Lemma succeq'_dec : rel_dec succeq'. 82 | 83 | Proof. 84 | intros l r. unfold succeq'. eapply dec. apply bcoef_pos_ok. 85 | Defined. 86 | 87 | Lemma succ'_sub : succ' << IR I succ. 88 | 89 | Proof. 90 | intros t u tu. unfold I, succ. set (r := mkRule t u). 91 | change t with (lhs r). change u with (rhs r). 92 | apply pi_compat_rule. hyp. 93 | Qed. 94 | 95 | Lemma succeq'_sub : succeq' << IR I succeq. 96 | 97 | Proof. 98 | intros t u tu. unfold I, succ. set (r := mkRule t u). 99 | change t with (lhs r). change u with (rhs r). 100 | apply pi_compat_rule_weak. hyp. 101 | Qed. 102 | 103 | Section ExtendedMonotoneAlgebra. 104 | 105 | Lemma monotone_succ : PolyStrongMonotone trsInt -> 106 | AWFMInterpretation.monotone I succ. 107 | 108 | Proof. 109 | unfold I. apply pi_monotone. 110 | Qed. 111 | 112 | End ExtendedMonotoneAlgebra. 113 | 114 | Import List. 115 | 116 | Section fin_Sig. 117 | 118 | Variable Fs : list Sig. 119 | Variable Fs_ok : forall f : Sig, In f Fs. 120 | 121 | Lemma fin_monotone_succ : 122 | forallb (fun f => bpstrong_monotone (trsInt f)) Fs = true -> 123 | AWFMInterpretation.monotone I succ. 124 | 125 | Proof. 126 | intro H. apply monotone_succ. intro f. rewrite <- bpstrong_monotone_ok. 127 | rewrite forallb_forall in H. apply H. apply Fs_ok. 128 | Qed. 129 | 130 | End fin_Sig. 131 | 132 | End MonotoneAlgebra. 133 | 134 | (***********************************************************************) 135 | (** tactics for Rainbow *) 136 | 137 | Arguments fin_monotone_succ [Fs] _ _ _ _ _ _ _ _ _ _ _. 138 | 139 | Ltac prove_cc_succ Fs_ok := 140 | apply IR_context_closed; apply (fin_monotone_succ Fs_ok); 141 | (check_eq || fail 10 "could not prove the monotony of this polynomial interpretation"). 142 | 143 | End PolyInt. 144 | -------------------------------------------------------------------------------- /SubtermCrit/ASimpleProj.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Sidi Ould-Biha & Frederic Blanqui, 2010-04-08 6 | 7 | simple projections for the subterm criterion 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import ATrs NatUtil BoolUtil VecUtil ListUtil LogicUtil 13 | RelUtil SN ACompat ARelation EqUtil. 14 | 15 | Section S. 16 | 17 | Variable Sig : Signature. 18 | 19 | Notation term := (term Sig). 20 | Notation rule := (rule Sig). Notation rules := (list rule). 21 | 22 | (***********************************************************************) 23 | (** projection definition *) 24 | 25 | Variable pi : Sig -> nat. 26 | 27 | Definition valid := forall f, arity f > 0 -> pi f < arity f. 28 | 29 | (***********************************************************************) 30 | (** projection function *) 31 | 32 | Variable Hpi : valid. 33 | 34 | Arguments Hpi [f] _. 35 | 36 | Definition proj t := 37 | match t with 38 | | Var _ => t 39 | | Fun f ts => 40 | match zerop (arity f) with 41 | | right H => Vnth ts (Hpi H) 42 | | left H => t 43 | end 44 | end. 45 | 46 | Definition proj_rule a := mkRule (proj (lhs a)) (proj (rhs a)). 47 | 48 | Definition proj_rules := map proj_rule. 49 | 50 | Lemma subterm_eq_proj : forall t, subterm_eq (proj t) t. 51 | 52 | Proof. 53 | intro. case t. simpl. intro. refl. 54 | simpl. intros. case (zerop (arity f)). intros. refl. 55 | intros. apply subterm_strict. apply subterm_fun. apply Vnth_in. 56 | Qed. 57 | 58 | (***********************************************************************) 59 | (** properties wrt substitutions *) 60 | 61 | Definition proj_sub s (x : variable) := proj (s x). 62 | 63 | Lemma proj_sub_var : forall s n, 64 | proj (sub s (Var n)) = sub (proj_sub s) (Var n). 65 | 66 | Proof. 67 | intros. simpl. auto. 68 | Qed. 69 | 70 | Lemma proj_sub_fun : forall s f ts, 71 | proj (sub s (Fun f ts)) = sub s (proj (Fun f ts)). 72 | 73 | Proof. 74 | intros. simpl. case (zerop (arity f)). intros. simpl. refl. 75 | intros. rewrite Vnth_map; auto. 76 | Qed. 77 | 78 | Lemma subterm_eq_proj_sub : forall s t, 79 | subterm_eq (proj (sub s t)) (sub s (proj t)). 80 | 81 | Proof. 82 | intros. case t. simpl. intros. apply subterm_eq_proj. 83 | intros. rewrite proj_sub_fun. refl. 84 | Qed. 85 | 86 | (***********************************************************************) 87 | (** properties preserved by projection *) 88 | 89 | Section proj_ord. 90 | 91 | Variable succ : relation term. 92 | 93 | Definition proj_ord : relation term := fun t u => succ (proj t) (proj u). 94 | 95 | (* preservation of transitivity *) 96 | Lemma proj_trans : transitive succ -> transitive proj_ord. 97 | 98 | Proof. 99 | intro. unfold transitive, proj_ord. intros. eapply H. apply H0. hyp. 100 | Qed. 101 | 102 | (* preservation of wellfoundedness *) 103 | Lemma WF_proj : WF succ -> WF proj_ord. 104 | 105 | Proof. 106 | intro. unfold proj_ord. apply (WF_inverse proj H). 107 | Qed. 108 | 109 | (* compatibility with rules *) 110 | Lemma compat_proj_ord : forall R : rules, 111 | compat succ (proj_rules R) -> compat proj_ord R. 112 | 113 | Proof. 114 | unfold compat. intros. unfold proj_ord. apply H. 115 | change (In (proj_rule (mkRule l r)) (proj_rules R)). 116 | apply in_map. hyp. 117 | Qed. 118 | 119 | End proj_ord. 120 | 121 | (* monotony wrt inclusion *) 122 | Lemma incl_proj : forall R S, R << S -> proj_ord R << proj_ord S. 123 | 124 | Proof. 125 | intros. unfold inclusion, proj_ord. intros. apply H. exact H0. 126 | Qed. 127 | 128 | (***********************************************************************) 129 | (** decision procedure for validity *) 130 | 131 | Variables (Fs : list Sig) (Fs_ok : forall f, In f Fs). 132 | 133 | Definition bvalid := 134 | forallb (fun f => beq_nat (arity f) 0 || bgt_nat (arity f) (pi f)) Fs. 135 | 136 | Lemma bvalid_ok : bvalid = true <-> valid. 137 | 138 | Proof. 139 | unfold valid, bvalid. apply forallb_ok_fintype. 2: hyp. 140 | intro f. rewrite orb_eq, beq_nat_ok, bgt_nat_ok. lia. 141 | Qed. 142 | 143 | End S. 144 | 145 | Arguments valid [Sig] _. 146 | Arguments bvalid [Sig] _ _. 147 | Arguments bvalid_ok [Sig] _ [Fs] _. 148 | Arguments proj [Sig pi] _ _. 149 | Arguments proj_ord [Sig pi] _ _ _ _. 150 | 151 | Ltac valid Fs_ok := rewrite <- (bvalid_ok _ Fs_ok); 152 | (check_eq || fail 10 "invalid simple projection"). 153 | -------------------------------------------------------------------------------- /Term/Varyadic/VContext.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-12-05 6 | 7 | one-hole contexts 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil. 13 | From CoLoR Require Import ListUtil. 14 | From CoLoR Require Export VTerm. 15 | 16 | Section S. 17 | 18 | Variable Sig : Signature. 19 | 20 | Notation term := (term Sig). Notation terms := (list term). 21 | 22 | (***********************************************************************) 23 | (** contexts and replacement of the hole *) 24 | 25 | Inductive context : Type := 26 | | Hole : context 27 | | Cont : forall f : Sig, terms -> context -> terms -> context. 28 | 29 | Fixpoint fill (c : context) (t : term) : term := 30 | match c with 31 | | Hole => t 32 | | Cont f v1 c' v2 => Fun f (v1 ++ fill c' t :: v2) 33 | end. 34 | 35 | Lemma var_eq_fill : forall x c t, Var x = fill c t -> c = Hole /\ t = Var x. 36 | 37 | Proof. intros. destruct c; simpl in H. auto. discr. Qed. 38 | 39 | (***********************************************************************) 40 | (** context composition *) 41 | 42 | Fixpoint comp (C : context) : context -> context := 43 | match C with 44 | | Hole => fun E => E 45 | | Cont f ts1 D ts2 => fun E => Cont f ts1 (comp D E) ts2 46 | end. 47 | 48 | Lemma fill_fill : forall C D u, fill C (fill D u) = fill (comp C D) u. 49 | 50 | Proof. induction C; simpl; intros. refl. rewrite (IHC D u). refl. Qed. 51 | 52 | (***********************************************************************) 53 | (** subterm ordering *) 54 | 55 | Definition subterm_eq u t := exists C, t = fill C u. 56 | 57 | Definition subterm u t := exists C, C <> Hole /\ t = fill C u. 58 | 59 | Lemma subterm_eq_var : forall u x, subterm_eq u (Var x) -> u = Var x. 60 | 61 | Proof. 62 | intros. unfold subterm_eq in H. destruct H as [C]. 63 | assert (C = Hole /\ u = Var x). 64 | apply var_eq_fill. exact H. destruct H0. exact H1. 65 | Qed. 66 | 67 | Lemma subterm_trans_eq2 : forall t u v, 68 | subterm t u -> subterm_eq u v -> subterm t v. 69 | 70 | Proof. 71 | unfold subterm, subterm_eq. intros. destruct H. destruct H. destruct H0. 72 | subst u. subst v. rewrite (fill_fill x0 x t). exists (comp x0 x). 73 | split. destruct x. cong. 74 | destruct x0; simpl; discr. refl. 75 | Qed. 76 | 77 | Lemma subterm_fun_elim : forall u f ts, 78 | subterm u (Fun f ts) -> exists t, In t ts /\ subterm_eq u t. 79 | 80 | Proof. 81 | intros. destruct H as [C [CH fC]]. 82 | destruct C. intuition. 83 | simpl in fC. injection fC. intros. subst ts. 84 | exists (fill C u). split. auto with datatypes. 85 | unfold subterm_eq. exists C. refl. 86 | Qed. 87 | 88 | Lemma subterm_immediate : forall f v a, In a v -> subterm a (Fun f v). 89 | 90 | Proof. 91 | intros. destruct (In_split a v) as [l1 [l2 dec]]. hyp. 92 | exists (Cont f l1 Hole l2). split. discr. simpl. congruence. 93 | Qed. 94 | 95 | (***********************************************************************) 96 | (** subterm-based induction principle *) 97 | 98 | Lemma subterm_eq_rect : forall (P : term -> Type) t, 99 | (forall u, subterm_eq u t -> P u) -> P t. 100 | 101 | Proof. intros. apply X. unfold subterm_eq. exists Hole. auto. Qed. 102 | 103 | Lemma subterm_sub_ind : forall (P : term -> Prop) 104 | (IH : forall t, (forall u, subterm u t -> P u) -> P t), 105 | forall t u, subterm_eq u t -> P u. 106 | 107 | Proof. 108 | intros P IH. set (P' := fun t => forall u, subterm_eq u t -> P u). 109 | change (forall t, P' t). apply term_ind_forall. 110 | (* var *) 111 | unfold P'. intros. assert (u = Var x). apply subterm_eq_var. hyp. 112 | subst u. apply IH. unfold subterm. intros. destruct H0. destruct H0. 113 | destruct x0. cong. discr. 114 | (* fun *) 115 | intros. unfold P'. intros. apply IH. intros. 116 | assert (subterm u0 (Fun f v)). eapply subterm_trans_eq2. apply H1. hyp. 117 | assert (exists t, In t v /\ subterm_eq u0 t). apply subterm_fun_elim with f. 118 | hyp. 119 | destruct H3. destruct H3. 120 | assert (P' x). apply lforall_in with term v. hyp. hyp. 121 | apply H5. hyp. 122 | Qed. 123 | 124 | Lemma subterm_ind : forall (P : term -> Prop) 125 | (IH : forall t, (forall u, subterm u t -> P u) -> P t), 126 | forall t, P t. 127 | 128 | Proof. intros. apply subterm_eq_rect. apply subterm_sub_ind. hyp. Qed. 129 | 130 | Lemma subterm_wf : well_founded subterm. 131 | 132 | Proof. 133 | intro t. induction t using subterm_ind. constructor. intros. 134 | apply H. hyp. 135 | Qed. 136 | 137 | End S. 138 | 139 | Arguments Hole {Sig}. 140 | -------------------------------------------------------------------------------- /Util/Set_/InfSet.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2014-12-11 6 | 7 | Infinite sets 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From Stdlib Require Import Basics Morphisms Setoid Lia. 13 | From CoLoR Require Import ClassicUtil IotaUtil EpsilonUtil LogicUtil SetUtil 14 | FinSet FunUtil BoundNat IotaUtil EpsilonUtil. 15 | 16 | Section S. 17 | 18 | Variable A : Type. 19 | 20 | (** * Definition of infinite sets *) 21 | 22 | Definition infinite (P : set A) := ~finite P. 23 | 24 | (** * Properties of [infinite]. *) 25 | 26 | Global Instance infinite_subset : Proper (subset ==> impl) infinite. 27 | 28 | Proof. intros P Q PQ P_inf Q_fin. rewrite <- PQ in Q_fin. contr. Qed. 29 | 30 | Global Instance infinite_equiv : Proper (equiv ==> iff) infinite. 31 | 32 | Proof. 33 | intros P Q PQ. unfold infinite. split. 34 | intros P_inf Q_fin. apply P_inf. rewrite PQ. hyp. 35 | intros Q_inf P_fin. apply Q_inf. rewrite <- PQ. hyp. 36 | Qed. 37 | 38 | (** A set [P] is infinite if there is an injection from [nat] to [P]. *) 39 | 40 | Lemma infinite_inj P : (exists f : nat -> elts P, injective f) -> infinite P. 41 | 42 | Proof. 43 | intros [f f_inj] [n [g [g_inj g_surj]]]. 44 | assert (S n <= n). 2: lia. 45 | apply N_inj_le with (f := inverse g_surj o f o (@N_val _)). 46 | inj. apply inj_N_val. 47 | Qed. 48 | 49 | (** * Set constructors preserve infiniteness. *) 50 | 51 | Lemma infinite_rem a P : infinite (rem a P) <-> infinite P. 52 | 53 | Proof. unfold infinite. gen (finite_rem_eq a P). tauto. Qed. 54 | 55 | Lemma infinite_union P Q : infinite (union P Q) <-> infinite P \/ infinite Q. 56 | 57 | Proof. unfold infinite. gen (finite_union_eq P Q). tauto. Qed. 58 | 59 | Lemma infinite_diff P Q : infinite (diff P Q) -> infinite P. 60 | 61 | Proof. intro P_Q_inf. intro P_fin. apply P_Q_inf. apply finite_diff. hyp. Qed. 62 | 63 | Lemma infinite_diff_finite P Q (Q_dec : forall x, {Q x}+{~Q x}) : 64 | infinite P -> finite Q -> infinite (diff P Q). 65 | 66 | Proof. 67 | intros P_inf Q_fin P_Q_fin. apply P_inf. cut (finite (union P Q)). 68 | rewrite finite_union_eq. fo. rewrite <- union_diff, finite_union_eq; auto. 69 | Qed. 70 | 71 | Lemma infinite_not_empty P : infinite P -> exists x, mem x P. 72 | 73 | Proof. 74 | intro P_inf. apply not_all_not_ex. intro h. 75 | assert (e : P [=] empty). fo. rewrite e in P_inf. 76 | gen finite_empty. fo. 77 | Qed. 78 | 79 | (****************************************************************************) 80 | (** * Type for infinite subsets of some set [P] *) 81 | 82 | Definition Pinf P := {Q | Q [<=] P /\ infinite Q}. 83 | 84 | Definition mk_Pinf P Q : Q [<=] P -> infinite Q -> Pinf P. 85 | 86 | Proof. intros QP Qinf. ex Q. auto. Defined. 87 | 88 | Definition Pinf_val P (Q : Pinf P) := proj1_sig Q. 89 | 90 | Global Coercion Pinf_val : Pinf >-> set. 91 | 92 | Definition Pinf_sub P (Q : Pinf P) : Q [<=] P. 93 | 94 | Proof. destruct Q as [Q [QP Qinf]]. hyp. Defined. 95 | 96 | Global Coercion Pinf_sub : Pinf >-> subset. 97 | 98 | Definition Pinf_inf P (Q : Pinf P) : infinite Q. 99 | 100 | Proof. destruct Q as [Q [QP Qinf]]. hyp. Defined. 101 | 102 | Global Coercion Pinf_inf : Pinf >-> infinite. 103 | 104 | Definition Pinf_subset P Q : P [<=] Q -> Pinf P -> Pinf Q. 105 | 106 | Proof. intros PQ [X [XP Xinf]]. ex X. split. trans P; hyp. hyp. Defined. 107 | 108 | Definition Pinf_self P (Q : Pinf P) : Pinf Q. 109 | 110 | Proof. ex Q. split. refl. destruct Q as [Q [QP Qinf]]; hyp. Defined. 111 | 112 | (****************************************************************************) 113 | (** An infinite set contains finite subsets of every cardinality. *) 114 | 115 | Section Pcard_of_inf. 116 | 117 | Variable W : set A. 118 | 119 | Lemma Pcard_of_inf_spec (P : Pinf W) : 120 | forall n, exists X : Pf A, X [<=] P /\ card X = S n. 121 | 122 | Proof. 123 | induction n. 124 | (* n=0 *) 125 | destruct (infinite_not_empty P) as [a aP]. ex (Pf_singleton a). 126 | split. 2: apply card_singleton. 127 | simpl. intro x. unfold impl. fo. subst. fo. 128 | (* n>0 *) 129 | destruct IHn as [X [XP cX]]. set (Q := diff P X). 130 | 131 | assert (Qinf : infinite Q). 132 | unfold Q. apply infinite_diff_finite. intro x. apply dec. 133 | destruct P as [P [Pall Pinf]]. hyp. 134 | destruct X as [X Xfin]; hyp. 135 | 136 | destruct (infinite_not_empty Qinf) as [a aQ]. ex (Pf_add a X). split. 137 | simpl. intro x. unfold impl. fo. subst. fo. 138 | rewrite card_add, cX. destruct (dec (mem a X)). fo. lia. 139 | Qed. 140 | 141 | Definition Pcard_of_inf (P : Pinf W) n : Pcard P (S n). 142 | 143 | Proof. 144 | destruct (cid (Pcard_of_inf_spec P n)) as [X [XP cX]]. ex X. auto. 145 | Defined. 146 | 147 | End Pcard_of_inf. 148 | 149 | End S. 150 | -------------------------------------------------------------------------------- /MannaNess/AMannaNess.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2006-12-01 6 | 7 | termination by using compatible reduction orderings 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import ATrs SN ARelation RelUtil ACompat LogicUtil. 13 | From Stdlib Require Import List. 14 | 15 | Section S. 16 | 17 | Variable Sig : Signature. 18 | 19 | Notation term := (term Sig). 20 | Notation rule := (rule Sig). Notation rules := (list rule). 21 | 22 | (***********************************************************************) 23 | (** Manna-Ness theorem (1970) *) 24 | 25 | Lemma manna_ness (R : rules) succ : 26 | reduction_ordering succ -> compat succ R -> WF (red R). 27 | 28 | Proof. 29 | intros. destruct H. apply (WF_incl succ). apply compat_red; auto. hyp. 30 | Qed. 31 | 32 | (***********************************************************************) 33 | (** an extension for proving the well-foundedness of relations of the form: 34 | several steps of R1 followed by a step of R2 *) 35 | 36 | Lemma manna_ness_mod (R E : rules) (rp : Reduction_pair Sig) : 37 | compat (rp_succ_eq rp) E -> compat (rp_succ rp) R -> WF (red_mod E R). 38 | 39 | Proof. 40 | intros. apply (WF_incl (rp_succ rp)). destruct rp; simpl in *. 41 | apply compat_red_mod with rp_succ_eq; try split; hyp. destruct rp; hyp. 42 | Qed. 43 | 44 | (***********************************************************************) 45 | (** an extension for proving the well-foundedness of relations of the form: 46 | several steps of R1 followed by a -head- step of R2 *) 47 | 48 | Lemma manna_ness_hd_mod (R E : rules) (wp : Weak_reduction_pair Sig) : 49 | compat (wp_succ_eq wp) E -> compat (wp_succ wp) R -> WF (hd_red_mod E R). 50 | 51 | Proof. 52 | intros. apply (WF_incl (wp_succ wp)). destruct wp; simpl in *. 53 | apply compat_hd_red_mod with wp_succ_eq; try split; hyp. destruct wp; hyp. 54 | Qed. 55 | 56 | (***********************************************************************) 57 | (** rule elimination *) 58 | 59 | Section rule_elimination. 60 | 61 | Variables R Rgt Rge : rules. 62 | 63 | Section mod. 64 | 65 | Variables E Egt Ege : rules. 66 | 67 | Lemma rule_elimination_mod : forall rp : Reduction_pair Sig, 68 | compat (rp_succ rp) Rgt -> 69 | compat (rp_succ_eq rp) Rge -> 70 | compat (rp_succ rp) Egt -> 71 | compat (rp_succ_eq rp) Ege -> 72 | WF (red_mod Ege Rge) -> 73 | WF (red_mod (Egt ++ Ege) (Rgt ++ Rge)). 74 | 75 | Proof with auto. 76 | intros. apply WF_incl with ((red Egt U red Ege)# @ (red Rgt U red Rge)). 77 | comp. apply rtc_incl. incl_trans (red Egt U red Ege)... 78 | apply red_union. 79 | apply red_union. 80 | apply wf_rel_mod... 81 | apply WF_incl with ((red (Rge ++ Ege)# @ (red (Rgt ++ Egt)))). 82 | comp. apply rtc_incl. apply red_union_inv. 83 | apply red_union_inv. 84 | apply manna_ness_mod with rp; apply compat_app... 85 | Qed. 86 | 87 | Lemma rule_elimination_hd_mod : forall wp : Weak_reduction_pair Sig, 88 | compat (wp_succ_eq wp) E -> 89 | compat (wp_succ_eq wp) Rge -> 90 | compat (wp_succ wp ) Rgt -> 91 | WF (hd_red_mod E Rge) -> 92 | WF (hd_red_mod E (Rgt ++ Rge)). 93 | 94 | Proof with auto. 95 | intros. apply WF_incl with (red E # @ (hd_red Rgt U hd_red Rge)). 96 | comp. apply hd_red_union. 97 | apply wf_rel_mod_simpl... 98 | apply WF_incl with (hd_red_mod (Rge ++ E) Rgt). 99 | comp. apply rtc_incl. 100 | incl_trans (red Rge U red E). union. apply hd_red_incl_red. 101 | apply red_union_inv. 102 | apply manna_ness_hd_mod with wp... 103 | apply compat_app... 104 | Qed. 105 | 106 | End mod. 107 | 108 | Lemma rule_elimination : forall rp : Reduction_pair Sig, 109 | compat (rp_succ_eq rp) Rge -> 110 | compat (rp_succ rp ) Rgt -> 111 | WF (red Rge) -> 112 | WF (red (Rgt ++ Rge)). 113 | 114 | Proof with auto. 115 | intros. eapply WF_incl. apply red_incl_red_mod. 116 | change (nil (A:=rule)) with (nil (A:=rule) ++ nil). 117 | apply rule_elimination_mod with rp... 118 | apply compat_empty. apply compat_empty. 119 | apply WF_incl with (red Rge)... apply red_mod_empty_incl_red. 120 | Qed. 121 | 122 | Lemma rule_elimination_hd_red : forall wp : Weak_reduction_pair Sig, 123 | compat (wp_succ_eq wp) Rge -> 124 | compat (wp_succ wp ) Rgt -> 125 | WF (hd_red Rge) -> 126 | WF (hd_red (Rgt ++ Rge)). 127 | 128 | Proof with auto. 129 | intros. eapply WF_incl. apply hd_red_incl_hd_red_mod. 130 | apply rule_elimination_hd_mod with wp... 131 | apply compat_empty. 132 | apply WF_incl with (hd_red Rge)... apply hd_red_mod_empty_incl_hd_red. 133 | Qed. 134 | 135 | End rule_elimination. 136 | 137 | End S. 138 | -------------------------------------------------------------------------------- /Coccinelle/basis/ordered_set.v: -------------------------------------------------------------------------------- 1 | (**************************************************************************) 2 | (* * *) 3 | (* _ * The Coccinelle Library / Evelyne Contejean *) 4 | (* * CNRS-LRI-Universite Paris Sud *) 5 | (* -/@|@\- * A3PAT Project *) 6 | (* -@ | @- * *) 7 | (* -\@|@/- * This file is distributed under the terms of the *) 8 | (* -v- * CeCILL-C licence *) 9 | (* * *) 10 | (**************************************************************************) 11 | 12 | From Stdlib Require Import Setoid Relations Arith FunInd. 13 | 14 | Inductive comp : Type := 15 | | Equivalent : comp 16 | | Less_than : comp 17 | | Greater_than : comp 18 | | Uncomparable : comp. 19 | 20 | Module Type S. 21 | 22 | Parameter A : Type. 23 | Parameter eq_bool : A -> A -> bool. 24 | Parameter eq_bool_ok : forall a1 a2, match eq_bool a1 a2 with true => a1 = a2 | false => ~ a1 = a2 end. 25 | 26 | Parameter o : relation A. 27 | Parameter o_bool : A -> A -> bool. 28 | Parameter o_bool_ok : forall a1 a2, match o_bool a1 a2 with true => o a1 a2 | false => ~ o a1 a2 end. 29 | Parameter o_proof : order A o. 30 | Parameter o_total : forall e1 e2 : A, {o e1 e2} + {o e2 e1}. 31 | 32 | End S. 33 | 34 | Module Type ES. 35 | 36 | Parameter A : Type. 37 | Parameter eq_A : relation A. 38 | Parameter eq_bool : A -> A -> bool. 39 | Parameter eq_bool_ok : forall a1 a2, match eq_bool a1 a2 with true => eq_A a1 a2 | false => ~eq_A a1 a2 end. 40 | Parameter eq_proof : equivalence A eq_A. 41 | 42 | Parameter o : relation A. 43 | Parameter o_compat : 44 | forall e1 e1' e2 e2', eq_A e1 e1' -> eq_A e2 e2' -> o e1 e2 -> o e1' e2'. 45 | Parameter o_bool : A -> A -> bool. 46 | Parameter o_bool_ok : forall a1 a2, match o_bool a1 a2 with true => o a1 a2 | false => ~ o a1 a2 end. 47 | Parameter o_proof : order A o. 48 | Parameter o_total : forall e1 e2 : A, {o e1 e2} + {o e2 e1}. 49 | 50 | Add Relation A eq_A 51 | reflexivity proved by (equiv_refl _ _ eq_proof) 52 | symmetry proved by (equiv_sym _ _ eq_proof) 53 | transitivity proved by (equiv_trans _ _ eq_proof) as EQA. 54 | 55 | End ES. 56 | 57 | (* 58 | Lemma toto : forall (A : Type) (R : relation A), 59 | (order A R) -> 60 | (forall e1 e2 : A, {R e1 e2} + {~ R e1 e2}) -> 61 | (forall e1 e2 : A, {R e1 e2} + {R e2 e1}) -> 62 | forall e1 e2 : A, {e1 = e2} + {e1 <> e2}. 63 | Proof. 64 | intros A R Ord Dec Tot e1 e2; destruct Ord. 65 | destruct (Tot e1 e2) as [e1_le_e2 | e2_le_e1]. 66 | destruct (Dec e2 e1) as [e2_le_e1 | not_e2_le_e1]. 67 | left; apply ord_antisym; trivial. 68 | right; intro; apply not_e2_le_e1; subst; trivial. 69 | destruct (Dec e1 e2) as [e1_le_e2 | not_e1_le_e2]. 70 | left; apply ord_antisym; trivial. 71 | right; intro; apply not_e1_le_e2; subst; trivial. 72 | Qed. 73 | *) 74 | 75 | Module Nat <: S with Definition A:=nat. 76 | 77 | Definition A := nat. 78 | Definition o : relation A := le. 79 | 80 | Lemma o_proof : order A o. 81 | Proof. 82 | unfold o; constructor; auto. 83 | unfold reflexive; auto. 84 | unfold transitive; intros n1 n2 n3 H1 H2; apply Nat.le_trans with n2; trivial. 85 | unfold antisymmetric; auto with arith. 86 | Qed. 87 | 88 | Lemma o_dec : forall e1 e2 : A, {o e1 e2} + {~ o e1 e2}. 89 | Proof. 90 | unfold o. 91 | induction e1; induction e2. 92 | left; trivial. 93 | left; auto with arith. 94 | right; auto with arith. 95 | destruct (IHe1 e2); destruct IHe2; auto with arith. 96 | Defined. 97 | 98 | Function eq_bool (n m : nat) {struct n} : bool := 99 | match n, m with 100 | | O, O => true 101 | | (S n'), (S m') => eq_bool n' m' 102 | | _, _ => false 103 | end. 104 | 105 | Lemma eq_bool_ok : forall a1 a2, match eq_bool a1 a2 with true => a1 = a2 | false => ~a1 = a2 end. 106 | fix eq_bool_ok 1. 107 | intros [ | n] [ | m]; simpl. 108 | reflexivity. 109 | discriminate. 110 | discriminate. 111 | generalize (eq_bool_ok n m); case (eq_bool n m). 112 | intros; subst; reflexivity. 113 | intros n_diff_m Sn_eq_Sm; apply n_diff_m; injection Sn_eq_Sm; intro; assumption. 114 | Defined. 115 | 116 | Function o_bool (n m : nat) {struct n} : bool := 117 | match n, m with 118 | | O, _ => true 119 | | (S n'), (S m') => o_bool n' m' 120 | | (S n'), O => false 121 | end. 122 | 123 | Lemma o_bool_ok : forall a1 a2, match o_bool a1 a2 with true => o a1 a2 | false => ~ o a1 a2 end. 124 | unfold o; fix o_bool_ok 1. 125 | intros [ | n] [ | m]; simpl. 126 | apply le_n. 127 | apply Nat.le_0_l. 128 | intro H; inversion H. 129 | generalize (o_bool_ok n m); case (o_bool n m). 130 | intro H; apply (le_n_S _ _ H). 131 | intros H H'; apply H; apply (le_S_n _ _ H'). 132 | Defined. 133 | 134 | Lemma o_total : forall e1 e2 : A, {o e1 e2} + {o e2 e1}. 135 | Proof. 136 | intros. 137 | unfold o. 138 | change ({e1<=e2}+{e1>=e2}). 139 | apply le_ge_dec. 140 | Qed. 141 | 142 | End Nat. 143 | 144 | -------------------------------------------------------------------------------- /RPO/VLPO.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Solange Coupet-Grimal and William Delobel, 2006-01-09 6 | 7 | Model of LPO statisfying Hypotheses in RPO_Types 8 | *) 9 | 10 | From CoLoR Require Import VPrecedence ListLex VRPO_Type VTerm ListUtil AccUtil 11 | LogicUtil. 12 | From Stdlib Require Import Relations. 13 | 14 | Module LPO (PT : VPrecedenceType). 15 | 16 | Module Export P := VPrecedence PT. 17 | Module Export S := Status PT. 18 | 19 | Inductive lt_lpo : relation term := 20 | | lpo1 : forall f g ss ts, g 21 | (forall t, In t ts -> lt_lpo t (Fun f ss)) -> 22 | lt_lpo (Fun g ts) (Fun f ss) 23 | | lpo2 : forall f g, f =F= g -> forall ss ts, 24 | (forall t, In t ts -> lt_lpo t (Fun f ss)) -> 25 | lex lt_lpo ts ss -> lt_lpo (Fun g ts) (Fun f ss) 26 | | lpo3 : forall t f ss, 27 | ex (fun s => In s ss /\ (s = t \/ lt_lpo t s)) -> 28 | lt_lpo t (Fun f ss). 29 | 30 | Definition mytau (f : Sig) (r : relation term) := lex r. 31 | 32 | End LPO. 33 | 34 | (***********************************************************************) 35 | 36 | From CoLoR Require Import VRPO_Type. 37 | 38 | Module LPO_Model (PT : VPrecedenceType) <: RPO_Model with Module P := PT. 39 | 40 | Module P := PT. 41 | Module Export LPO := LPO PT. 42 | 43 | Definition tau := mytau. 44 | 45 | Lemma tau_dec : forall f R ts ss, 46 | (forall t s, In t ts -> In s ss -> {R t s} + {~R t s}) -> 47 | {tau f R ts ss} + {~tau f R ts ss}. 48 | 49 | Proof. 50 | unfold tau, mytau. intros. apply lex_status_dec. hyp. 51 | Defined. 52 | 53 | Lemma status_eq : forall f g, f =F= g -> tau f = tau g. 54 | 55 | Proof. 56 | auto. 57 | Qed. 58 | 59 | Definition lt := lt_lpo. 60 | 61 | (* Verification des axiomes lt : *) 62 | 63 | Lemma lt_roots : forall f g, g forall ss ts, 64 | (forall t, In t ts -> lt t (Fun f ss)) -> lt (Fun g ts) (Fun f ss). 65 | 66 | Proof. 67 | intros f g ltgf ss ts Hsub. unfold lt; apply lpo1; trivial. 68 | Qed. 69 | 70 | Lemma lt_status : forall f g, f =F= g -> 71 | forall ss ts, (forall t, In t ts -> lt t (Fun f ss)) -> 72 | tau f lt ts ss -> lt (Fun g ts) (Fun f ss). 73 | 74 | Proof. 75 | unfold lt, tau, mytau. intros f g feqg ss ts Hsub. apply lpo2; trivial. 76 | Qed. 77 | 78 | Lemma lt_subterm : forall f ss t, 79 | ex (fun s => In s ss /\ (s = t \/ lt t s)) -> lt t (Fun f ss). 80 | 81 | Proof. 82 | intros f ss t Hex; unfold lt; apply lpo3; trivial. 83 | Qed. 84 | 85 | Lemma lt_decomp : forall s t, lt s t -> 86 | ((ex (fun f => ex (fun g => ex (fun ss => ex (fun ts => 87 | ltF f g /\ 88 | (forall s, In s ss -> lt s (Fun g ts)) /\ 89 | s = Fun f ss /\ 90 | t = Fun g ts 91 | ))))) 92 | \/ 93 | ex (fun f => ex (fun g => ex (fun ss => ex (fun ts => 94 | f =F= g /\ 95 | (forall s, In s ss -> lt s (Fun g ts)) /\ 96 | tau f lt ss ts /\ 97 | s = Fun f ss /\ 98 | t = Fun g ts))))) 99 | \/ 100 | ex (fun f => ex (fun ts => ex (fun t' => 101 | In t' ts /\ (s = t' \/ lt s t')) /\ t = Fun f ts)). 102 | 103 | Proof. 104 | intros s t Hst. 105 | unfold lt in Hst; inversion Hst as [f'' g'' ss' ts'' ltFfg Hsi' 106 | | f'' g'' feq ss' ts'' Hsub Hss'ts' Hsi' 107 | | f'' s' ts'' Hex]; subst. 108 | left; left. 109 | exists g''; exists f''; exists ts''; exists ss'. 110 | split; trivial; repeat split; trivial. 111 | left; right. 112 | exists g''; exists f''; exists ts''; exists ss'. 113 | split; elim feq; trivial; repeat split; trivial. 114 | right. 115 | exists s'; exists ts''. 116 | split; trivial; elim Hex; clear Hex; intros t Ht. 117 | elim Ht; clear Ht; intros t_in_ts Ht. 118 | exists t; split; trivial. 119 | elim Ht; [left; subst | right]; trivial. 120 | Qed. 121 | 122 | (***********************************************************************) 123 | 124 | Lemma SPO_to_status_SPO : forall f (r : relation term), 125 | forall ss, 126 | (forall s, In s ss -> 127 | (forall t u, r s t -> r t u -> r s u) /\ (r s s -> False)) -> 128 | (forall ts us, tau f r ss ts -> tau f r ts us -> tau f r ss us) 129 | /\ (tau f r ss ss -> False). 130 | 131 | Proof. 132 | intros f r. unfold tau, mytau. apply SPO_to_lex_SPO. 133 | Qed. 134 | 135 | Lemma mono_axiom : forall f (r : relation term), 136 | forall ss ts, one_less r ss ts -> tau f r ss ts. 137 | 138 | Proof. 139 | intros f r ss ts. unfold tau, mytau. apply one_less2lex. 140 | Qed. 141 | 142 | (***********************************************************************) 143 | 144 | Definition wf_ltF := ltF_wf. 145 | Definition leF_dec := leF_dec. 146 | 147 | Definition lifting R := forall l, Accs lt l -> Restricted_acc (Accs lt) R l. 148 | 149 | Lemma status_lifting : forall f, lifting (tau f lt). 150 | 151 | Proof. 152 | intro f; unfold tau, mytau, lt, lifting. apply (lex_lifting lt_lpo). 153 | Qed. 154 | 155 | End LPO_Model. 156 | -------------------------------------------------------------------------------- /Term/WithArity/ASignature.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Sebastien Hinderer, 2004-02-09 6 | 7 | signature for algebraic terms with arity 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import EqUtil ListUtil LogicUtil. 13 | 14 | (***********************************************************************) 15 | (** Variables are represented by natural numbers. *) 16 | 17 | Notation variable := nat (only parsing). 18 | 19 | (***********************************************************************) 20 | (** Signature with a decidable set of symbols of fixed arity. *) 21 | 22 | Record Signature : Type := mkSignature { 23 | symbol :> Type; 24 | arity : symbol -> nat; 25 | beq_symb : symbol -> symbol -> bool; 26 | beq_symb_ok : forall x y, beq_symb x y = true <-> x = y 27 | }. 28 | 29 | Arguments mkSignature [symbol] _ [beq_symb] _. 30 | Arguments arity [s] _. 31 | Arguments beq_symb [s] _ _. 32 | Arguments beq_symb_ok {s x y}. 33 | 34 | Ltac case_beq_symb Sig := EqUtil.case_beq (@beq_symb Sig) (@beq_symb_ok Sig). 35 | 36 | Definition eq_symb_dec Sig := dec_beq (@beq_symb_ok Sig). 37 | 38 | Arguments eq_symb_dec [Sig] _ _. 39 | 40 | (***********************************************************************) 41 | (** Tactic for proving beq_symb_ok *) 42 | 43 | Ltac beq_symb_ok := intros f g; split; 44 | [destruct f; destruct g; simpl; intro; (discr || refl) 45 | | intro; subst g; destruct f; refl]. 46 | 47 | (***********************************************************************) 48 | (** Modules and module types for signatures *) 49 | 50 | Module Type SIG. 51 | Parameter Sig : Signature. 52 | End SIG. 53 | 54 | (* DecidableType *) 55 | From Stdlib Require Import DecidableType. 56 | Module DecType (Import S : SIG) <: DecidableType. 57 | Definition t := @symbol Sig. 58 | Definition eq := @eq t. 59 | Definition eq_refl := @eq_refl t. 60 | Definition eq_sym := @eq_sym t. 61 | Definition eq_trans := @eq_trans t. 62 | Definition eq_dec := EqUtil.dec_beq (@beq_symb_ok Sig). 63 | End DecType. 64 | 65 | (* finite signature *) 66 | Module Type FSIG. 67 | Parameter Sig : Signature. 68 | Parameter Fs : list Sig. 69 | Parameter Fs_ok : forall f, In f Fs. 70 | End FSIG. 71 | 72 | (***********************************************************************) 73 | (** Weighted signatures *) 74 | 75 | Module Type WSIG. 76 | Parameter Sig : Signature. 77 | Parameter weight : Sig -> nat. 78 | Parameter weight_inj : forall f g, weight f = weight g -> f = g. 79 | End WSIG. 80 | 81 | From CoLoR Require Import NatUtil BoolUtil. 82 | 83 | Section weight_inj. 84 | 85 | Variable Sig : Signature. 86 | Variable Fs : list Sig. 87 | Variable Fs_ok : forall f, In f Fs. 88 | Variable weight : Sig -> nat. 89 | 90 | Definition bweight_inj := 91 | forallb (fun f => forallb (fun g => 92 | implb (beq_nat (weight f) (weight g)) (beq_symb f g)) Fs) Fs. 93 | 94 | Lemma bweight_inj_ok : bweight_inj = true <-> 95 | (forall f g, weight f = weight g -> f = g). 96 | 97 | Proof. 98 | unfold bweight_inj. apply forallb_ok_fintype. 2: hyp. intro f. 99 | apply forallb_ok_fintype. 2: hyp. intro g. unfold implb. 100 | case_eq (beq_nat (weight f) (weight g)). 101 | rewrite beq_nat_ok, beq_symb_ok. tauto. 102 | rewrite (beq_ko beq_nat_ok). tauto. 103 | Qed. 104 | 105 | End weight_inj. 106 | 107 | (*Arguments bweight_inj_ok [Sig Fs] _ [weight].*) 108 | 109 | Ltac weight_inj Fs_ok := rewrite <- (bweight_inj_ok Fs_ok); 110 | (check_eq || fail 10 "non-injective weight function"). 111 | 112 | (***********************************************************************) 113 | (** Ordered signatures *) 114 | 115 | From Stdlib Require Import OrderedType. 116 | 117 | Module OrdType (Import S : WSIG) <: OrderedType. 118 | 119 | Include (DecType S). 120 | 121 | Definition lt f g := weight f < weight g. 122 | 123 | Lemma lt_trans : forall x y z, lt x y -> lt y z -> lt x z. 124 | 125 | Proof. unfold lt. intros. lia. Qed. 126 | 127 | Lemma lt_not_eq : forall x y, lt x y -> x <> y. 128 | 129 | Proof. unfold lt. intros x y l e. subst. lia. Qed. 130 | 131 | Definition blt f g := bgt_nat (weight g) (weight f). 132 | 133 | Lemma blt_ok : forall f g, blt f g = true <-> lt f g. 134 | 135 | Proof. intros f g. unfold blt, lt. apply bgt_nat_ok. Qed. 136 | 137 | Lemma lt_total : forall f g, 138 | beq_symb f g = false -> blt f g = false -> blt g f = true. 139 | 140 | Proof. 141 | intros f g. unfold blt. 142 | rewrite bgt_nat_ok, bgt_nat_ko, (beq_ko (@beq_symb_ok Sig)). intros. 143 | destruct (eq_nat_dec (weight f) (weight g)). 144 | ded (weight_inj e). subst g. cong. 145 | lia. 146 | Qed. 147 | 148 | Definition compare : forall x y, Compare lt (@Logic.eq t) x y. 149 | 150 | Proof. 151 | intros x y. case_eq (beq_symb x y); intro H. 152 | rewrite beq_symb_ok in H. apply (EQ H). 153 | case_eq (blt x y); intro H0. 154 | rewrite blt_ok in H0. apply (LT H0). 155 | eapply GT. rewrite <- blt_ok. apply lt_total; hyp. 156 | Defined. 157 | 158 | End OrdType. 159 | -------------------------------------------------------------------------------- /Term/Varyadic/VSignature.v: -------------------------------------------------------------------------------- 1 | (** 2 | CoLoR, a Coq library on rewriting and termination. 3 | See the COPYRIGHTS and LICENSE files. 4 | 5 | - Frederic Blanqui, 2005-06-10 6 | 7 | signature for algebraic terms with no arity 8 | *) 9 | 10 | Set Implicit Arguments. 11 | 12 | From CoLoR Require Import LogicUtil. 13 | 14 | (***********************************************************************) 15 | (** Variables are represented by natural numbers. *) 16 | 17 | Notation variable := nat (only parsing). 18 | 19 | (***********************************************************************) 20 | (** Signature with a decidable set of symbols of fixed arity. *) 21 | 22 | Record Signature : Type := mkSignature { 23 | symbol :> Type; 24 | beq_symb : symbol -> symbol -> bool; 25 | beq_symb_ok : forall x y, beq_symb x y = true <-> x = y 26 | }. 27 | 28 | Arguments mkSignature [symbol beq_symb] _. 29 | Arguments beq_symb [s] _ _. 30 | Arguments beq_symb_ok {s x y}. 31 | 32 | From CoLoR Require Import EqUtil. 33 | 34 | Ltac case_beq_symb Sig := EqUtil.case_beq (@beq_symb Sig) (@beq_symb_ok Sig). 35 | 36 | Definition eq_symb_dec Sig := dec_beq (@beq_symb_ok Sig). 37 | 38 | Arguments eq_symb_dec [Sig] _ _. 39 | 40 | (***********************************************************************) 41 | (** Tactic for proving beq_symb_ok *) 42 | 43 | Ltac beq_symb_ok := intros f g; split; 44 | [destruct f; destruct g; simpl; intro; (discr || refl) 45 | | intro; subst g; destruct f; refl]. 46 | 47 | (***********************************************************************) 48 | (** Modules and module types for signatures *) 49 | 50 | Module Type SIG. 51 | Parameter Sig : Signature. 52 | End SIG. 53 | 54 | (* DecidableType *) 55 | From Stdlib Require Import DecidableType. 56 | Module DecType (Import S : SIG) <: DecidableType. 57 | Definition t := @symbol Sig. 58 | Definition eq := @eq t. 59 | Definition eq_refl := @eq_refl t. 60 | Definition eq_sym := @eq_sym t. 61 | Definition eq_trans := @eq_trans t. 62 | Definition eq_dec := EqUtil.dec_beq (@beq_symb_ok Sig). 63 | End DecType. 64 | 65 | (* finite signature *) 66 | From CoLoR Require Import ListUtil. 67 | Module Type FSIG. 68 | Parameter Sig : Signature. 69 | Parameter Fs : list Sig. 70 | Parameter Fs_ok : forall f, In f Fs. 71 | End FSIG. 72 | 73 | (***********************************************************************) 74 | (** Weighted signatures *) 75 | 76 | Module Type WSIG. 77 | Parameter Sig : Signature. 78 | Parameter weight : Sig -> nat. 79 | Parameter weight_inj : forall f g, weight f = weight g -> f = g. 80 | End WSIG. 81 | 82 | From CoLoR Require Import NatUtil BoolUtil. 83 | 84 | Section weight_inj. 85 | 86 | Variable Sig : Signature. 87 | Variable Fs : list Sig. 88 | Variable Fs_ok : forall f, In f Fs. 89 | Variable weight : Sig -> nat. 90 | 91 | Definition bweight_inj := 92 | forallb (fun f => forallb (fun g => 93 | implb (beq_nat (weight f) (weight g)) (beq_symb f g)) Fs) Fs. 94 | 95 | Lemma bweight_inj_ok : bweight_inj = true <-> 96 | (forall f g, weight f = weight g -> f = g). 97 | 98 | Proof. 99 | unfold bweight_inj. apply forallb_ok_fintype. 2: hyp. intro f. 100 | apply forallb_ok_fintype. 2: hyp. intro g. unfold implb. 101 | case_eq (beq_nat (weight f) (weight g)). 102 | rewrite beq_nat_ok, beq_symb_ok. tauto. 103 | rewrite (beq_ko beq_nat_ok). tauto. 104 | Qed. 105 | 106 | End weight_inj. 107 | 108 | (*Arguments bweight_inj_ok _ [Fs] _ [weight].*) 109 | 110 | Ltac weight_inj Fs_ok := rewrite <- (bweight_inj_ok _ Fs_ok); 111 | (check_eq || fail 10 "non-injective weight function"). 112 | 113 | (***********************************************************************) 114 | (** Ordered signatures *) 115 | 116 | From Stdlib Require Import OrderedType. 117 | 118 | Module OrdType (Import S : WSIG) <: OrderedType. 119 | 120 | Include (DecType S). 121 | 122 | Definition lt f g := weight f < weight g. 123 | 124 | Lemma lt_trans : forall x y z, lt x y -> lt y z -> lt x z. 125 | 126 | Proof. unfold lt. intros. lia. Qed. 127 | 128 | Lemma lt_not_eq : forall x y, lt x y -> x <> y. 129 | 130 | Proof. unfold lt. intros x y l e. subst. lia. Qed. 131 | 132 | Definition blt f g := bgt_nat (weight g) (weight f). 133 | 134 | Lemma blt_ok : forall f g, blt f g = true <-> lt f g. 135 | 136 | Proof. intros f g. unfold blt, lt. apply bgt_nat_ok. Qed. 137 | 138 | Lemma lt_total : forall f g, 139 | beq_symb f g = false -> blt f g = false -> blt g f = true. 140 | 141 | Proof. 142 | intros f g. unfold blt. 143 | rewrite bgt_nat_ok, bgt_nat_ko, (beq_ko (@beq_symb_ok Sig)). intros. 144 | destruct (eq_nat_dec (weight f) (weight g)). 145 | ded (weight_inj e). subst g. cong. 146 | lia. 147 | Qed. 148 | 149 | Definition compare : forall x y, Compare lt (@Logic.eq t) x y. 150 | 151 | Proof. 152 | intros x y. case_eq (beq_symb x y); intro H. 153 | rewrite beq_symb_ok in H. apply (EQ H). 154 | case_eq (blt x y); intro H0. 155 | rewrite blt_ok in H0. apply (LT H0). 156 | eapply GT. rewrite <- blt_ok. apply lt_total; hyp. 157 | Defined. 158 | 159 | End OrdType. 160 | --------------------------------------------------------------------------------