├── .gitignore ├── BlocksWorldExample1.v ├── Context.v ├── LinLog.v ├── LinProp.v ├── Makefile ├── PermutationHelpers.v ├── README.md ├── Sig.v └── _RocqProject /.gitignore: -------------------------------------------------------------------------------- 1 | *~ 2 | *.bak 3 | 4 | .coqdeps.d 5 | Makefile 6 | Makefile.conf 7 | .Makefile.d 8 | 9 | *.o 10 | *.v.crashcoqide 11 | *.v.d 12 | *.vio 13 | *.vo 14 | *.vok 15 | *.vos 16 | *.glob 17 | .*.aux 18 | .coq-native/ 19 | .csdp.cache 20 | .lia.cache 21 | .nia.cache 22 | .nlia.cache 23 | .nra.cache 24 | csdp.cache 25 | lia.cache 26 | nia.cache 27 | nlia.cache 28 | nra.cache 29 | -------------------------------------------------------------------------------- /BlocksWorldExample1.v: -------------------------------------------------------------------------------- 1 | From Stdlib Require Import List. 2 | Import ListNotations. 3 | 4 | (** Basic table-top blocks manipulation demonstration. *) 5 | Parameter (Block:Type). 6 | 7 | (** First define a type to capture the language used to describe the 8 | world. These are the propositions that apply to our domain. *) 9 | Module BlockProp. 10 | Inductive BlockProp : Type := 11 | | on : Block -> Block -> BlockProp 12 | | table : Block -> BlockProp 13 | | clear : Block -> BlockProp 14 | | holds : BlockProp 15 | | empty : BlockProp. 16 | End BlockProp. 17 | 18 | Require Import Sig. 19 | Module BlockSig <: Sig. 20 | Definition A := BlockProp.BlockProp. 21 | End BlockSig. 22 | 23 | (** Instantiate the linear logic connectives over our langauge. *) 24 | Require Context. 25 | Module BlocksWorld := Context.Context BlockSig. 26 | Import BlocksWorld. 27 | 28 | (** Define linear logic propositions for the terms of our language. *) 29 | Definition on X Y : LinProp := Term (BlockProp.on X Y). 30 | Definition table X : LinProp := Term (BlockProp.table X). 31 | Definition clear X : LinProp := Term (BlockProp.clear X). 32 | Definition holds : LinProp := Term BlockProp.holds. 33 | Definition empty : LinProp := Term BlockProp.empty. 34 | 35 | (** Define the actions our system is capable of. *) 36 | Axiom newBlock : [empty] ⊢ holds. 37 | Axiom put : forall X Y, [holds] ⊢ empty ⊗ clear X ⊗ (table X & clear Y ⊸ on X Y). 38 | Axiom get : forall X Y, 39 | [empty, clear X] ⊢ holds ⊗ (table X ⊸ One & on X Y ⊸ clear Y). 40 | 41 | (** An example table-top manipulation task. *) 42 | 43 | (** Start with an empty gripper, and block [X] on top of block [Y]. We 44 | take block [X], put it on the table, then get block [Z] and put it 45 | on top of block [Y]. *) 46 | Theorem test1 : forall X Y Z, [empty, clear X, on X Y] ⊢ on Z Y ⊗ ⊤. 47 | Proof. intros. 48 | break_context_at 2. 49 | apply Cut with (B:=holds⊗(table X ⊸ One & on X Y ⊸ clear Y)). 50 | apply get. 51 | apply context_app_comm. 52 | apply Cut with (B:=holds⊗(on X Y ⊸ clear Y)). 53 | apply product_as_context. 54 | break_context_at 1. 55 | timesIntro; [|withElimR]; refl. 56 | apply product_app_as_context. simpl. 57 | match goal with 58 | |- [?A,?B,?C] ⊢ _ => permute_context [A,C,B] 59 | end. 60 | break_context_at 2. 61 | apply Cut with (B:=clear Y). 62 | break_context_at 1; apply context_app_comm. 63 | impliesElim; refl. 64 | apply Cut with (B:=empty ⊗ clear Z ⊗ (table Z & clear Y ⊸ on Z Y)). 65 | apply put. product_to_context. 66 | permute_context ([table Z & clear Y ⊸ on Z Y]++[clear Y, empty, clear Z]). 67 | apply Cut with (B:=clear Y ⊸ on Z Y). 68 | withElimR; refl. 69 | simpl. 70 | permute_context ([clear Y ⊸ on Z Y, clear Y]++[empty, clear Z]). 71 | apply Cut with (B:=on Z Y). 72 | break_context_at 1. 73 | impliesElim; refl. 74 | simpl. 75 | permute_context ([on Z Y] ++ [empty, clear Z]). 76 | timesIntro. refl. topIntro. 77 | Qed. 78 | -------------------------------------------------------------------------------- /Context.v: -------------------------------------------------------------------------------- 1 | (** This module layers context handling facilities on top of the core 2 | linear logic module, [LinLog]. The most commonly useful definitions 3 | in this module are the tactics [product_to_context], 4 | [break_context_at], and [permute_context]. *) 5 | Require Import Sig. 6 | Require LinLog. 7 | Require Export PermutationHelpers. 8 | Module Context(Term:Sig). 9 | Module LinLogTerm := LinLog.LinLog Term. 10 | Export LinLogTerm. 11 | Import List Permutation. 12 | Import ListNotations. 13 | 14 | Definition Context := list LinProp. 15 | 16 | (** Convert a product into a list. *) 17 | Fixpoint product_to_context (P:LinProp) : Context := 18 | match P with 19 | | x ⊗ y => x :: product_to_context y 20 | | x => [x] 21 | end. 22 | 23 | (** If we take ⊤ as a right-identity on products, we can convert a 24 | list of propositions into a single product. *) 25 | Fixpoint context_to_product (C:Context) : LinProp := 26 | match C with 27 | | [] => ⊤ 28 | | [x] => x 29 | | h::t => h ⊗ context_to_product t 30 | end. 31 | 32 | Lemma context_to_product_left_inverse : 33 | forall P, context_to_product (product_to_context P) = P. 34 | Proof. induction P; auto. simpl. rewrite IHP2. 35 | destruct P2; auto. 36 | Qed. 37 | 38 | (** A product consisting of the two elements of a context is just as 39 | good as the context represented as a list. *) 40 | Lemma product_as_context : forall A B C, [A,B] ⊢ C -> [A ⊗ B] ⊢ C. 41 | Proof. intros. apply Elim. replace [A⊗B] with ([] ++ [A⊗B]) by auto. 42 | apply TimesElim with (A:=A) (B:=B). refl. assumption. 43 | Qed. 44 | 45 | Lemma product_app_as_context : forall A B C D, 46 | A++[B,C] ⊢ D -> A ++ [B⊗C] ⊢ D. 47 | Proof. 48 | intros. apply Elim. apply TimesElim with (A:=B) (B:=C). refl. assumption. 49 | Qed. 50 | 51 | (** Rewrite a list append expression to pop the had of the right 52 | operand onto the tail of the left operand. *) 53 | Lemma app_pop_cons : forall {A} P (X:A) T, P++(X::T) = (P++[X])++T. 54 | Proof. intros. replace (X::T) with ([X]++T) by auto. apply app_assoc. Qed. 55 | 56 | Ltac simplify_all_apps := 57 | match goal with 58 | | |- context f [(?P ++ ?X::?Y)++?Z] => 59 | let tmp := fresh in 60 | assert ((P++X::Y)++Z = P++(X::Y++Z)) as tmp by 61 | (rewrite <- app_assoc; simpl; reflexivity); 62 | simpl in tmp; rewrite tmp; clear tmp 63 | | _ => idtac 64 | end. 65 | 66 | (** Flatten a product into a [Context]. *) 67 | Ltac product_to_context := 68 | match goal with 69 | | |- [_⊗_] ⊢ _ => apply product_as_context; 70 | match goal with 71 | | |- ?P::?X ⊢ _ => replace (P::X) with ([P]++X) by auto; 72 | product_to_context 73 | end 74 | | |- ?P++[_⊗_] ⊢ _ => apply product_app_as_context; 75 | rewrite app_pop_cons; 76 | product_to_context 77 | | _ => simpl; repeat simplify_all_apps 78 | end. 79 | 80 | (** [break_at n lst] returns a pair whose first component is no more 81 | than the first [n] elements of [list], and whose second component 82 | is the remaining elements of [lst]. *) 83 | Definition break_at {A} (n:nat) (L:list A) := 84 | let fix go n pre l := 85 | match n with 86 | | 0 => (rev pre, l) 87 | | S n' => match l with 88 | | [] => (rev pre, []) 89 | | h::t => go n' (h::pre) t 90 | end 91 | end in 92 | go n [] L. 93 | 94 | (** Tactic for splitting a context into two appended sub-lists. *) 95 | Ltac break_context_at i := 96 | match goal with 97 | |- ?C ⊢ _ => let parts := fresh in 98 | set (parts:=break_at i C); simpl in parts; 99 | let pre := fresh in 100 | set (pre:=fst parts); simpl in pre; 101 | let pos := fresh in 102 | set (pos:=snd parts); simpl in pos; 103 | let hperm := fresh in 104 | assert (hperm:Permutation C (pre++pos)) by auto; 105 | rewrite hperm; 106 | clear parts; subst pre; subst pos; clear hperm 107 | end. 108 | 109 | Lemma context_app_comm : forall A B C, A++B ⊢ C -> B++A ⊢ C. 110 | Proof. intros. 111 | assert (hperm:Permutation (A++B) (B++A)) by apply Permutation_app_comm. 112 | rewrite <- hperm. assumption. 113 | Qed. 114 | 115 | (** Tactic for permuting a context. If the goal is [Context ⊢ 116 | Conclusion], then [permute_context Context'] will attempt to 117 | replace [Context] with [Context'] in the goal. *) 118 | Ltac permute_context Desired := 119 | match goal with 120 | |- ?C ⊢ _ => let H := fresh in 121 | assert (H:Permutation C Desired) by permut; 122 | rewrite H; clear H 123 | end. 124 | 125 | End Context. 126 | -------------------------------------------------------------------------------- /LinLog.v: -------------------------------------------------------------------------------- 1 | Require Import Sig. 2 | From Stdlib Require Export List. 3 | From Stdlib Require Export Permutation. 4 | From Stdlib Require Import Setoid. 5 | 6 | Require LinProp. 7 | Module LinLog(Term:Sig). 8 | Module LinPropTerm := LinProp.LinProp Term. 9 | Import Term. 10 | Export LinPropTerm. 11 | 12 | Notation "[]" := nil. 13 | Notation "[ x , .. , y ]" := (x :: .. (y :: nil) ..) : list_scope. 14 | 15 | (** The linear consequence operator is defined inductively between any 16 | list of linear propositions (the hypotheses) and some other 17 | proposition (the conclusion). *) 18 | 19 | Reserved Notation "A ⊢ B" (at level 74, right associativity). 20 | Reserved Notation "A ⊢ι B" (at level 74, right associativity). 21 | Reserved Notation "A ⊢ε B" (at level 74, right associativity). 22 | 23 | (** Define the linear logic consequence relation by partitioning it 24 | into introduction and elimination forms. *) 25 | Inductive LinCons : list LinProp -> LinProp -> Prop := 26 | | Intro : forall Γ C, Γ ⊢ι C -> Γ ⊢ C 27 | | Elim : forall Γ C, Γ ⊢ε C -> Γ ⊢ C 28 | | PermContext : forall Γ Δ C, Permutation Γ Δ -> Γ ⊢ C -> Δ ⊢ C 29 | | Cut : forall Γ Δ A B, Γ ⊢ B -> Δ ++ [B] ⊢ A -> Γ ++ Δ ⊢ A 30 | where "Γ ⊢ C" := (LinCons Γ C) 31 | with 32 | LinIntro : list LinProp -> LinProp -> Prop := 33 | | PermContextι : forall Γ Δ C, Permutation Γ Δ -> Γ ⊢ι C -> Δ ⊢ι C 34 | | Refl : forall t, [t] ⊢ι t 35 | | OneIntro : [] ⊢ι One 36 | | TopIntro : forall Γ, Γ ⊢ι ⊤ 37 | | TimesIntro : forall Δ Θ A B, Δ ⊢ A -> Θ ⊢ B -> Δ ++ Θ ⊢ι A ⊗ B 38 | | ImpliesIntro : forall Γ A B, Γ ++ [A] ⊢ B -> Γ ⊢ι A ⊸ B 39 | | WithIntro : forall Γ A B, Γ ⊢ A -> Γ ⊢ B -> Γ ⊢ι A & B 40 | | PlusIntroL : forall Γ A B, Γ ⊢ A -> Γ ⊢ι A ⊕ B 41 | | PlusIntroR : forall Γ A B, Γ ⊢ B -> Γ ⊢ι A ⊕ B 42 | where "Γ ⊢ι C" := (LinIntro Γ C) 43 | with 44 | LinElim : list LinProp -> LinProp -> Prop := 45 | | PermContextε : forall Γ Δ C, Permutation Γ Δ -> Γ ⊢ε C -> Δ ⊢ε C 46 | | OneElim : forall Δ Θ A, Δ ⊢ One -> Θ ⊢ A -> Δ ++ Θ ⊢ε A 47 | | TimesElim : forall Γ Δ A B C, Δ ⊢ A ⊗ B -> Γ ++ [A,B] ⊢ C -> Γ ++ Δ ⊢ε C 48 | | ImpliesElim : forall Γ Δ A B, Γ ⊢ A ⊸ B -> Δ ⊢ A -> Γ ++ Δ ⊢ε B 49 | | WithElimL : forall Γ A B, Γ ⊢ A & B -> Γ ⊢ε A 50 | | WithElimR : forall Γ A B, Γ ⊢ A & B -> Γ ⊢ε B 51 | | PlusElim : forall Θ Δ A B C, 52 | Θ ⊢ A ⊕ B -> Δ ++ [A] ⊢ C -> Δ ++ [B] ⊢ C -> Θ ++ Δ ⊢ε C 53 | | ZeroElim : forall Γ A, Γ ++ [⊥] ⊢ε A 54 | where "Γ ⊢ε C" := (LinElim Γ C). 55 | 56 | Arguments PermContext [Γ Δ C]. 57 | Arguments Cut [Γ Δ A B]. 58 | 59 | Create HintDb linlog. 60 | #[local] Hint Constructors LinIntro LinElim : linlog. 61 | 62 | Add Parametric Morphism : LinIntro with 63 | signature (@Permutation LinProp) ==> Logic.eq ==> iff as LinIntro_morph. 64 | intros X Y HPerm C. split; intros H. inversion H; subst; try (eauto with linlog). 65 | symmetry in HPerm. apply PermContextι with (Δ:=X) in H; assumption. 66 | Qed. 67 | 68 | Add Parametric Morphism : LinElim with 69 | signature (@Permutation LinProp) ==> Logic.eq ==> iff as LinElim_morph. 70 | intros X Y HPerm C. split; intros H. inversion H; subst; try (eauto with linlog). 71 | symmetry in HPerm. apply PermContextε with (Δ:=X) in H; assumption. 72 | Qed. 73 | 74 | Add Parametric Morphism : LinCons with 75 | signature (@Permutation LinProp) ==> Logic.eq ==> iff as LinCons_morph. 76 | intros X Y HPerm C. split; intros H; inversion H. 77 | rewrite HPerm in H0. apply Intro. assumption. 78 | rewrite HPerm in H0. apply Elim. assumption. 79 | apply (PermContext HPerm). assumption. 80 | apply (Cut H0) in H1. rewrite <- H2 in HPerm. 81 | apply (PermContext HPerm). assumption. 82 | 83 | symmetry in HPerm. rewrite HPerm in H0. apply Intro. assumption. 84 | symmetry in HPerm. rewrite HPerm in H0. apply Elim. assumption. 85 | symmetry in HPerm. apply (PermContext HPerm). assumption. 86 | symmetry in HPerm. apply (Cut H0) in H1. rewrite <- H2 in HPerm. 87 | apply (PermContext HPerm). assumption. 88 | Qed. 89 | 90 | (** * Tactics that de-clutter constructor application. *) 91 | 92 | Ltac refl := apply Intro; apply Refl. 93 | Ltac topIntro := apply Intro; apply TopIntro. 94 | Ltac withElimL := apply Elim; eapply WithElimL. 95 | Ltac withElimR := apply Elim; eapply WithElimR. 96 | Ltac withIntro := apply Intro; apply WithIntro. 97 | Ltac impliesElim := apply Elim; eapply ImpliesElim. 98 | Ltac timesIntro := apply Intro; apply TimesIntro. 99 | 100 | End LinLog. -------------------------------------------------------------------------------- /LinProp.v: -------------------------------------------------------------------------------- 1 | (** Define the linear logic connectives. This module is parameterized 2 | over a term type that may be embedded in linear logic propositions. *) 3 | Require Import Sig. 4 | Module LinProp(Term:Sig). 5 | Import Term. 6 | Inductive LinProp : Type := 7 | | Implies : LinProp -> LinProp -> LinProp 8 | | One : LinProp 9 | | Plus : LinProp -> LinProp -> LinProp 10 | | Times : LinProp -> LinProp -> LinProp 11 | | Top : LinProp 12 | | With : LinProp -> LinProp -> LinProp 13 | | Zero : LinProp 14 | | Term : A -> LinProp. 15 | 16 | Infix "⊗" := Times (at level 41, right associativity). (* \otimes *) 17 | Infix "&" := With (at level 74, right associativity). 18 | Infix "⊕" := Plus (at level 51, right associativity). (* \oplus *) 19 | Notation "⊤" := Top. (* \top *) 20 | Notation "⊥" := Zero. (* \bot *) 21 | Infix "⊸" := Implies (at level 73, right associativity). (* \multimap *) 22 | End LinProp. 23 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | ########################################################################## 2 | ## # The Rocq Prover / The Rocq Development Team ## 3 | ## v # Copyright INRIA, CNRS and contributors ## 4 | ## /dev/null 2>/dev/null; echo $$?)) 82 | STDTIME?=command time -f $(TIMEFMT) 83 | else 84 | ifeq (0,$(shell gtime -f "" true >/dev/null 2>/dev/null; echo $$?)) 85 | STDTIME?=gtime -f $(TIMEFMT) 86 | else 87 | STDTIME?=command time 88 | endif 89 | endif 90 | else 91 | STDTIME?=command time -f $(TIMEFMT) 92 | endif 93 | 94 | COQBIN?= 95 | ifneq (,$(COQBIN)) 96 | # add an ending / 97 | COQBIN:=$(COQBIN)/ 98 | endif 99 | 100 | # Coq binaries 101 | ROCQ ?= "$(COQBIN)rocq" 102 | COQC ?= "$(COQBIN)rocq" c 103 | COQTOP ?= "$(COQBIN)rocq" repl 104 | COQCHK ?= "$(COQBIN)rocqchk" 105 | COQNATIVE ?= "$(COQBIN)rocq" native-precompile 106 | COQDEP ?= "$(COQBIN)rocq" dep 107 | COQDOC ?= "$(COQBIN)rocq" doc 108 | COQPP ?= "$(COQBIN)rocq" pp-mlg 109 | COQMKFILE ?= "$(COQBIN)rocq" makefile 110 | OCAMLLIBDEP ?= "$(COQBIN)ocamllibdep" 111 | 112 | # Timing scripts 113 | COQMAKE_ONE_TIME_FILE ?= "$(COQCORELIB)/tools/make-one-time-file.py" 114 | COQMAKE_BOTH_TIME_FILES ?= "$(COQCORELIB)/tools/make-both-time-files.py" 115 | COQMAKE_BOTH_SINGLE_TIMING_FILES ?= "$(COQCORELIB)/tools/make-both-single-timing-files.py" 116 | BEFORE ?= 117 | AFTER ?= 118 | 119 | # OCaml binaries 120 | CAMLC ?= "$(OCAMLFIND)" ocamlc -c 121 | CAMLOPTC ?= "$(OCAMLFIND)" opt -c 122 | CAMLLINK ?= "$(OCAMLFIND)" ocamlc -linkall 123 | CAMLOPTLINK ?= "$(OCAMLFIND)" opt -linkall 124 | CAMLDOC ?= "$(OCAMLFIND)" ocamldoc 125 | CAMLDEP ?= "$(OCAMLFIND)" ocamldep -slash -ml-synonym .mlpack 126 | 127 | # DESTDIR is prepended to all installation paths 128 | DESTDIR ?= 129 | 130 | # Debug builds, typically -g to OCaml, -debug to Rocq. 131 | CAMLDEBUG ?= 132 | COQDEBUG ?= 133 | 134 | # Extra packages to be linked in (as in findlib -package) 135 | CAMLPKGS ?= 136 | FINDLIBPKGS = -package rocq-runtime.plugins.ltac $(CAMLPKGS) 137 | 138 | # Option for making timing files 139 | TIMING?= 140 | # Option for changing sorting of timing output file 141 | TIMING_SORT_BY ?= auto 142 | # Option for changing the fuzz parameter on the output file 143 | TIMING_FUZZ ?= 0 144 | # Option for changing whether to use real or user time for timing tables 145 | TIMING_REAL?= 146 | # Option for including the memory column(s) 147 | TIMING_INCLUDE_MEM?= 148 | # Option for sorting by the memory column 149 | TIMING_SORT_BY_MEM?= 150 | # Output file names for timed builds 151 | TIME_OF_BUILD_FILE ?= time-of-build.log 152 | TIME_OF_BUILD_BEFORE_FILE ?= time-of-build-before.log 153 | TIME_OF_BUILD_AFTER_FILE ?= time-of-build-after.log 154 | TIME_OF_PRETTY_BUILD_FILE ?= time-of-build-pretty.log 155 | TIME_OF_PRETTY_BOTH_BUILD_FILE ?= time-of-build-both.log 156 | TIME_OF_PRETTY_BUILD_EXTRA_FILES ?= - # also output to the command line 157 | 158 | TGTS ?= 159 | 160 | # Retro compatibility (DESTDIR is standard on Unix, DSTROOT is not) 161 | ifdef DSTROOT 162 | DESTDIR := $(DSTROOT) 163 | endif 164 | 165 | # Substitution of the path by appending $(DESTDIR) if needed. 166 | # The variable $(COQMF_WINDRIVE) can be needed for Cygwin environments. 167 | windrive_path = $(if $(COQMF_WINDRIVE),$(subst $(COQMF_WINDRIVE),/,$(1)),$(1)) 168 | destination_path = $(if $(DESTDIR),$(DESTDIR)/$(call windrive_path,$(1)),$(1)) 169 | 170 | # Installation paths of libraries and documentation. 171 | COQLIBINSTALL ?= $(call destination_path,$(COQLIB)/user-contrib) 172 | COQDOCINSTALL ?= $(call destination_path,$(DOCDIR)/coq/user-contrib) 173 | COQPLUGININSTALL ?= $(call destination_path,$(COQCORELIB)/..) 174 | COQTOPINSTALL ?= $(call destination_path,$(COQLIB)/toploop) # FIXME: Unused variable? 175 | 176 | # findlib files installation 177 | FINDLIBPREINST= mkdir -p "$(COQPLUGININSTALL)/" 178 | FINDLIBDESTDIR= -destdir "$(COQPLUGININSTALL)/" 179 | 180 | # we need to move out of sight $(METAFILE) otherwise findlib thinks the 181 | # package is already installed 182 | findlib_install = \ 183 | $(HIDE)if [ "$(METAFILE)" ]; then \ 184 | $(FINDLIBPREINST) && \ 185 | mv "$(METAFILE)" "$(METAFILE).skip" ; \ 186 | "$(OCAMLFIND)" install $(2) $(FINDLIBDESTDIR) $(FINDLIBPACKAGE) $(1); \ 187 | rc=$$?; \ 188 | mv "$(METAFILE).skip" "$(METAFILE)"; \ 189 | exit $$rc; \ 190 | fi 191 | findlib_remove = \ 192 | $(HIDE)if [ ! -z "$(METAFILE)" ]; then\ 193 | "$(OCAMLFIND)" remove $(FINDLIBDESTDIR) $(FINDLIBPACKAGE); \ 194 | fi 195 | 196 | 197 | ########## End of parameters ################################################## 198 | # What follows may be relevant to you only if you need to 199 | # extend this Makefile. If so, look for 'Extension point' here and 200 | # put in Makefile.local double colon rules accordingly. 201 | # E.g. to perform some work after the all target completes you can write 202 | # 203 | # post-all:: 204 | # echo "All done!" 205 | # 206 | # in Makefile.local 207 | # 208 | ############################################################################### 209 | 210 | 211 | 212 | 213 | # Flags ####################################################################### 214 | # 215 | # We define a bunch of variables combining the parameters. 216 | # To add additional flags to coq, coqchk or coqdoc, set the 217 | # {COQ,COQCHK,COQDOC}EXTRAFLAGS variable to whatever you want to add. 218 | # To overwrite the default choice and set your own flags entirely, set the 219 | # {COQ,COQCHK,COQDOC}FLAGS variable. 220 | 221 | SHOW := $(if $(VERBOSE),@true "",@echo "") 222 | HIDE := $(if $(VERBOSE),,@) 223 | 224 | TIMER=$(if $(TIMED), $(STDTIME), $(TIMECMD)) 225 | 226 | OPT?= 227 | 228 | # The DYNLIB variable is used by "coqdep -dyndep var" in .v.d 229 | ifeq '$(OPT)' '-byte' 230 | USEBYTE:=true 231 | DYNLIB:=.cma 232 | else 233 | USEBYTE:= 234 | DYNLIB:=.cmxs 235 | endif 236 | 237 | # these variables are meant to be overridden if you want to add *extra* flags 238 | COQEXTRAFLAGS?= 239 | COQCHKEXTRAFLAGS?= 240 | COQDOCEXTRAFLAGS?= 241 | 242 | # Find the last argument of the form "-native-compiler FLAG" 243 | COQUSERNATIVEFLAG:=$(strip \ 244 | $(subst -native-compiler-,,\ 245 | $(lastword \ 246 | $(filter -native-compiler-%,\ 247 | $(subst -native-compiler ,-native-compiler-,\ 248 | $(strip $(COQEXTRAFLAGS))))))) 249 | 250 | COQFILTEREDEXTRAFLAGS:=$(strip \ 251 | $(filter-out -native-compiler-%,\ 252 | $(subst -native-compiler ,-native-compiler-,\ 253 | $(strip $(COQEXTRAFLAGS))))) 254 | 255 | COQACTUALNATIVEFLAG:=$(lastword $(COQMF_COQ_NATIVE_COMPILER_DEFAULT) $(COQMF_COQPROJECTNATIVEFLAG) $(COQUSERNATIVEFLAG)) 256 | 257 | ifeq '$(COQACTUALNATIVEFLAG)' 'yes' 258 | COQNATIVEFLAG="-w" "-deprecated-native-compiler-option" "-native-compiler" "ondemand" 259 | COQDONATIVE="yes" 260 | else 261 | ifeq '$(COQACTUALNATIVEFLAG)' 'ondemand' 262 | COQNATIVEFLAG="-w" "-deprecated-native-compiler-option" "-native-compiler" "ondemand" 263 | COQDONATIVE="no" 264 | else 265 | COQNATIVEFLAG="-w" "-deprecated-native-compiler-option" "-native-compiler" "no" 266 | COQDONATIVE="no" 267 | endif 268 | endif 269 | 270 | # these flags do NOT contain the libraries, to make them easier to overwrite 271 | COQFLAGS?=-q $(OTHERFLAGS) $(COQFILTEREDEXTRAFLAGS) $(COQNATIVEFLAG) 272 | COQCHKFLAGS?=-silent -o $(COQCHKEXTRAFLAGS) 273 | COQDOCFLAGS?=-interpolate -utf8 $(COQDOCEXTRAFLAGS) 274 | 275 | COQDOCLIBS?=$(COQLIBS_NOML) 276 | 277 | # The version of Coq being run and the version of rocq makefile that 278 | # generated this makefile 279 | # NB --print-version is not in the rocq shim 280 | COQ_VERSION:=$(shell $(ROCQ) c --print-version | cut -d " " -f 1) 281 | COQMAKEFILE_VERSION:=9.0.0 282 | 283 | # COQ_SRC_SUBDIRS is for user-overriding, usually to add 284 | # `user-contrib/Foo` to the includes, we keep COQCORE_SRC_SUBDIRS for 285 | # Coq's own core libraries, which should be replaced by ocamlfind 286 | # options at some point. 287 | COQ_SRC_SUBDIRS?= 288 | COQSRCLIBS?= $(foreach d,$(COQ_SRC_SUBDIRS), -I "$(COQLIB)/$(d)") 289 | 290 | CAMLFLAGS+=$(OCAMLLIBS) $(COQSRCLIBS) 291 | # ocamldoc fails with unknown argument otherwise 292 | CAMLDOCFLAGS:=$(filter-out -annot, $(filter-out -bin-annot, $(CAMLFLAGS))) 293 | CAMLFLAGS+=$(OCAMLWARN) 294 | 295 | ifneq (,$(TIMING)) 296 | ifeq (after,$(TIMING)) 297 | TIMING_EXT=after-timing 298 | else 299 | ifeq (before,$(TIMING)) 300 | TIMING_EXT=before-timing 301 | else 302 | TIMING_EXT=timing 303 | endif 304 | endif 305 | TIMING_ARG=-time-file $<.$(TIMING_EXT) 306 | else 307 | TIMING_ARG= 308 | endif 309 | 310 | ifneq (,$(PROFILING)) 311 | PROFILE_ARG=-profile $@.prof.json 312 | PROFILE_ZIP=gzip -f $@.prof.json 313 | else 314 | PROFILE_ARG= 315 | PROFILE_ZIP=true 316 | endif 317 | 318 | # Files ####################################################################### 319 | # 320 | # We here define a bunch of variables about the files being part of the 321 | # Rocq project in order to ease the writing of build target and build rules 322 | 323 | VDFILE := .Makefile.d 324 | 325 | ALLSRCFILES := \ 326 | $(MLGFILES) \ 327 | $(MLFILES) \ 328 | $(MLPACKFILES) \ 329 | $(MLLIBFILES) \ 330 | $(MLIFILES) 331 | 332 | # helpers 333 | vo_to_obj = $(addsuffix .o,\ 334 | $(filter-out Warning: Error:,\ 335 | $(shell $(COQTOP) -q -noinit -batch -quiet -print-mod-uid $(1)))) 336 | strip_dotslash = $(patsubst ./%,%,$(1)) 337 | 338 | # without this we get undefined variables in the expansion for the 339 | # targets of the [deprecated,use-mllib-or-mlpack] rule 340 | with_undef = $(if $(filter-out undefined, $(origin $(1))),$($(1))) 341 | 342 | VO = vo 343 | VOS = vos 344 | 345 | VOFILES = $(VFILES:.v=.$(VO)) 346 | GLOBFILES = $(VFILES:.v=.glob) 347 | HTMLFILES = $(VFILES:.v=.html) 348 | GHTMLFILES = $(VFILES:.v=.g.html) 349 | BEAUTYFILES = $(addsuffix .beautified,$(VFILES)) 350 | TEXFILES = $(VFILES:.v=.tex) 351 | GTEXFILES = $(VFILES:.v=.g.tex) 352 | CMOFILES = \ 353 | $(MLGFILES:.mlg=.cmo) \ 354 | $(MLFILES:.ml=.cmo) \ 355 | $(MLPACKFILES:.mlpack=.cmo) 356 | CMXFILES = $(CMOFILES:.cmo=.cmx) 357 | OFILES = $(CMXFILES:.cmx=.o) 358 | CMAFILES = $(MLLIBFILES:.mllib=.cma) $(MLPACKFILES:.mlpack=.cma) 359 | CMXAFILES = $(CMAFILES:.cma=.cmxa) 360 | CMIFILES = \ 361 | $(CMOFILES:.cmo=.cmi) \ 362 | $(MLIFILES:.mli=.cmi) 363 | # the /if/ is because old _CoqProject did not list a .ml(pack|lib) but just 364 | # a .mlg file 365 | CMXSFILES = \ 366 | $(MLPACKFILES:.mlpack=.cmxs) \ 367 | $(CMXAFILES:.cmxa=.cmxs) \ 368 | $(if $(MLPACKFILES)$(CMXAFILES),,\ 369 | $(MLGFILES:.mlg=.cmxs) $(MLFILES:.ml=.cmxs)) 370 | 371 | # files that are packed into a plugin (no extension) 372 | PACKEDFILES = \ 373 | $(call strip_dotslash, \ 374 | $(foreach lib, \ 375 | $(call strip_dotslash, \ 376 | $(MLPACKFILES:.mlpack=_MLPACK_DEPENDENCIES)),$(call with_undef,$(lib)))) 377 | # files that are archived into a .cma (mllib) 378 | LIBEDFILES = \ 379 | $(call strip_dotslash, \ 380 | $(foreach lib, \ 381 | $(call strip_dotslash, \ 382 | $(MLLIBFILES:.mllib=_MLLIB_DEPENDENCIES)),$(call with_undef,$(lib)))) 383 | CMIFILESTOINSTALL = $(filter-out $(addsuffix .cmi,$(PACKEDFILES)),$(CMIFILES)) 384 | CMOFILESTOINSTALL = $(filter-out $(addsuffix .cmo,$(PACKEDFILES)),$(CMOFILES)) 385 | OBJFILES = $(call vo_to_obj,$(VOFILES)) 386 | ALLNATIVEFILES = \ 387 | $(OBJFILES:.o=.cmi) \ 388 | $(OBJFILES:.o=.cmx) \ 389 | $(OBJFILES:.o=.cmxs) 390 | FINDLIBPACKAGE=$(patsubst .%,%,$(suffix $(METAFILE))) 391 | 392 | # trick: wildcard filters out non-existing files, so that `install` doesn't show 393 | # warnings and `clean` doesn't pass to rm a list of files that is too long for 394 | # the shell. 395 | NATIVEFILES = $(wildcard $(ALLNATIVEFILES)) 396 | FILESTOINSTALL = \ 397 | $(VOFILES) \ 398 | $(VFILES) \ 399 | $(GLOBFILES) \ 400 | $(NATIVEFILES) 401 | FINDLIBFILESTOINSTALL = \ 402 | $(CMIFILESTOINSTALL) 403 | ifeq '$(HASNATDYNLINK)' 'true' 404 | DO_NATDYNLINK = yes 405 | FINDLIBFILESTOINSTALL += $(CMXSFILES) $(CMXAFILES) $(CMOFILESTOINSTALL:.cmo=.cmx) 406 | else 407 | DO_NATDYNLINK = 408 | endif 409 | 410 | ALLDFILES = $(addsuffix .d,$(ALLSRCFILES)) $(VDFILE) 411 | 412 | # Compilation targets ######################################################### 413 | 414 | all: 415 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all 416 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all 417 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all 418 | .PHONY: all 419 | 420 | all.timing.diff: 421 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all 422 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all.timing.diff TIME_OF_PRETTY_BUILD_EXTRA_FILES="" 423 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all 424 | .PHONY: all.timing.diff 425 | 426 | ifeq (0,$(TIMING_REAL)) 427 | TIMING_REAL_ARG := 428 | TIMING_USER_ARG := --user 429 | else 430 | ifeq (1,$(TIMING_REAL)) 431 | TIMING_REAL_ARG := --real 432 | TIMING_USER_ARG := 433 | else 434 | TIMING_REAL_ARG := 435 | TIMING_USER_ARG := 436 | endif 437 | endif 438 | 439 | ifeq (0,$(TIMING_INCLUDE_MEM)) 440 | TIMING_INCLUDE_MEM_ARG := --no-include-mem 441 | else 442 | TIMING_INCLUDE_MEM_ARG := 443 | endif 444 | 445 | ifeq (1,$(TIMING_SORT_BY_MEM)) 446 | TIMING_SORT_BY_MEM_ARG := --sort-by-mem 447 | else 448 | TIMING_SORT_BY_MEM_ARG := 449 | endif 450 | 451 | make-pretty-timed-before:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_BEFORE_FILE) 452 | make-pretty-timed-after:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_AFTER_FILE) 453 | make-pretty-timed make-pretty-timed-before make-pretty-timed-after:: 454 | $(HIDE)rm -f pretty-timed-success.ok 455 | $(HIDE)($(MAKE) --no-print-directory -f "$(PARENT)" $(TGTS) TIMED=1 2>&1 && touch pretty-timed-success.ok) | tee -a $(TIME_OF_BUILD_FILE) 456 | $(HIDE)rm pretty-timed-success.ok # must not be -f; must fail if the touch failed 457 | print-pretty-timed:: 458 | $(HIDE)$(COQMAKE_ONE_TIME_FILE) $(TIMING_INCLUDE_MEM_ARG) $(TIMING_SORT_BY_MEM_ARG) $(TIMING_REAL_ARG) $(TIME_OF_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 459 | print-pretty-timed-diff:: 460 | $(HIDE)$(COQMAKE_BOTH_TIME_FILES) --sort-by=$(TIMING_SORT_BY) $(TIMING_INCLUDE_MEM_ARG) $(TIMING_SORT_BY_MEM_ARG) $(TIMING_REAL_ARG) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 461 | ifeq (,$(BEFORE)) 462 | print-pretty-single-time-diff:: 463 | @echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing' 464 | $(HIDE)false 465 | else 466 | ifeq (,$(AFTER)) 467 | print-pretty-single-time-diff:: 468 | @echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing' 469 | $(HIDE)false 470 | else 471 | print-pretty-single-time-diff:: 472 | $(HIDE)$(COQMAKE_BOTH_SINGLE_TIMING_FILES) --fuzz=$(TIMING_FUZZ) --sort-by=$(TIMING_SORT_BY) $(TIMING_USER_ARG) $(AFTER) $(BEFORE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 473 | endif 474 | endif 475 | pretty-timed: 476 | $(HIDE)$(MAKE) --no-print-directory -f "$(PARENT)" make-pretty-timed 477 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-timed 478 | .PHONY: pretty-timed make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff 479 | 480 | # Extension points for actions to be performed before/after the all target 481 | pre-all:: 482 | @# Extension point 483 | $(HIDE)if [ "$(COQMAKEFILE_VERSION)" != "$(COQ_VERSION)" ]; then\ 484 | echo "W: This Makefile was generated by Rocq/Coq $(COQMAKEFILE_VERSION)";\ 485 | echo "W: while the current Rocq version is $(COQ_VERSION)";\ 486 | fi 487 | .PHONY: pre-all 488 | 489 | post-all:: 490 | @# Extension point 491 | .PHONY: post-all 492 | 493 | real-all: $(VOFILES) $(if $(USEBYTE),bytefiles,optfiles) 494 | .PHONY: real-all 495 | 496 | real-all.timing.diff: $(VOFILES:.vo=.v.timing.diff) 497 | .PHONY: real-all.timing.diff 498 | 499 | bytefiles: $(CMOFILES) $(CMAFILES) 500 | .PHONY: bytefiles 501 | 502 | optfiles: $(if $(DO_NATDYNLINK),$(CMXSFILES)) 503 | .PHONY: optfiles 504 | 505 | vos: $(VOFILES:%.vo=%.vos) 506 | .PHONY: vos 507 | 508 | vok: $(VOFILES:%.vo=%.vok) 509 | .PHONY: vok 510 | 511 | validate: $(VOFILES) 512 | $(TIMER) $(COQCHK) $(COQCHKFLAGS) $(COQLIBS_NOML) $(PROFILE_ARG) $^ 513 | $(HIDE)$(PROFILE_ZIP) 514 | .PHONY: validate 515 | 516 | only: $(TGTS) 517 | .PHONY: only 518 | 519 | # Documentation targets ####################################################### 520 | 521 | html: $(GLOBFILES) $(VFILES) 522 | $(SHOW)'COQDOC -d html $(GAL)' 523 | $(HIDE)mkdir -p html 524 | $(HIDE)$(COQDOC) \ 525 | -toc $(COQDOCFLAGS) -html $(GAL) $(COQDOCLIBS) -d html $(VFILES) 526 | 527 | mlihtml: $(MLIFILES:.mli=.cmi) 528 | $(SHOW)'CAMLDOC -d $@' 529 | $(HIDE)mkdir $@ || rm -rf $@/* 530 | $(HIDE)$(CAMLDOC) -html \ 531 | -d $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES) $(FINDLIBPKGS) 532 | 533 | all-mli.tex: $(MLIFILES:.mli=.cmi) 534 | $(SHOW)'CAMLDOC -latex $@' 535 | $(HIDE)$(CAMLDOC) -latex \ 536 | -o $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES) $(FINDLIBPKGS) 537 | 538 | all.ps: $(VFILES) 539 | $(SHOW)'COQDOC -ps $(GAL)' 540 | $(HIDE)$(COQDOC) \ 541 | -toc $(COQDOCFLAGS) -ps $(GAL) $(COQDOCLIBS) \ 542 | -o $@ `$(COQDEP) -sort $(VFILES)` 543 | 544 | all.pdf: $(VFILES) 545 | $(SHOW)'COQDOC -pdf $(GAL)' 546 | $(HIDE)$(COQDOC) \ 547 | -toc $(COQDOCFLAGS) -pdf $(GAL) $(COQDOCLIBS) \ 548 | -o $@ `$(COQDEP) -sort $(VFILES)` 549 | 550 | # FIXME: not quite right, since the output name is different 551 | gallinahtml: GAL=-g 552 | gallinahtml: html 553 | 554 | all-gal.ps: GAL=-g 555 | all-gal.ps: all.ps 556 | 557 | all-gal.pdf: GAL=-g 558 | all-gal.pdf: all.pdf 559 | 560 | # ? 561 | beautify: $(BEAUTYFILES) 562 | for file in $^; do mv $${file%.beautified} $${file%beautified}old && mv $${file} $${file%.beautified}; done 563 | @echo 'Do not do "make clean" until you are sure that everything went well!' 564 | @echo 'If there were a problem, execute "for file in $$(find . -name \*.v.old -print); do mv $${file} $${file%.old}; done" in your shell/' 565 | .PHONY: beautify 566 | 567 | # Installation targets ######################################################## 568 | # 569 | # There rules can be extended in Makefile.local 570 | # Extensions can't assume when they run. 571 | 572 | # We use $(file) to avoid generating a very long command string to pass to the shell 573 | # (cf https://coq.zulipchat.com/#narrow/stream/250632-Coq-Platform-devs-.26-users/topic/Strange.20command.20length.20limit.20on.20Linux) 574 | # However Apple ships old make which doesn't have $(file) so we need a fallback 575 | $(file >.hasfile,1) 576 | HASFILE:=$(shell if [ -e .hasfile ]; then echo 1; rm .hasfile; fi) 577 | 578 | MKFILESTOINSTALL= $(if $(HASFILE),$(file >.filestoinstall,$(FILESTOINSTALL)),\ 579 | $(shell rm -f .filestoinstall) \ 580 | $(foreach x,$(FILESTOINSTALL),$(shell printf '%s\n' "$x" >> .filestoinstall))) 581 | 582 | # findlib needs the package to not be installed, so we remove it before 583 | # installing it (see the call to findlib_remove) 584 | install: META 585 | @$(MKFILESTOINSTALL) 586 | $(HIDE)code=0; for f in $$(cat .filestoinstall); do\ 587 | if ! [ -f "$$f" ]; then >&2 echo $$f does not exist; code=1; fi \ 588 | done; exit $$code 589 | $(HIDE)for f in $$(cat .filestoinstall); do\ 590 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\ 591 | if [ "$$?" != "0" -o -z "$$df" ]; then\ 592 | echo SKIP "$$f" since it has no logical path;\ 593 | else\ 594 | install -d "$(COQLIBINSTALL)/$$df" &&\ 595 | install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\ 596 | echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\ 597 | fi;\ 598 | done 599 | $(call findlib_remove) 600 | $(call findlib_install, META $(FINDLIBFILESTOINSTALL)) 601 | $(HIDE)$(MAKE) install-extra -f "$(SELF)" 602 | @rm -f .filestoinstall 603 | install-extra:: 604 | @# Extension point 605 | .PHONY: install install-extra 606 | 607 | META: $(METAFILE) 608 | $(HIDE)if [ "$(METAFILE)" ]; then \ 609 | cat "$(METAFILE)" | grep -v 'directory.*=.*' > META; \ 610 | fi 611 | 612 | install-byte: 613 | $(call findlib_install, $(CMAFILES) $(CMOFILESTOINSTALL), -add) 614 | 615 | install-doc:: html mlihtml 616 | @# Extension point 617 | $(HIDE)install -d "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" 618 | $(HIDE)for i in html/*; do \ 619 | dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\ 620 | install -m 0644 "$$i" "$$dest";\ 621 | echo INSTALL "$$i" "$$dest";\ 622 | done 623 | $(HIDE)install -d \ 624 | "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml" 625 | $(HIDE)for i in mlihtml/*; do \ 626 | dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\ 627 | install -m 0644 "$$i" "$$dest";\ 628 | echo INSTALL "$$i" "$$dest";\ 629 | done 630 | .PHONY: install-doc 631 | 632 | uninstall:: 633 | @# Extension point 634 | @$(MKFILESTOINSTALL) 635 | $(call findlib_remove) 636 | $(HIDE)for f in $$(cat .filestoinstall); do \ 637 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`" &&\ 638 | instf="$(COQLIBINSTALL)/$$df/`basename $$f`" &&\ 639 | rm -f "$$instf" &&\ 640 | echo RM "$$instf" ;\ 641 | done 642 | $(HIDE)for f in $$(cat .filestoinstall); do \ 643 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`" &&\ 644 | echo RMDIR "$(COQLIBINSTALL)/$$df/" &&\ 645 | (rmdir "$(COQLIBINSTALL)/$$df/" 2>/dev/null || true); \ 646 | done 647 | @rm -f .filestoinstall 648 | 649 | .PHONY: uninstall 650 | 651 | uninstall-doc:: 652 | @# Extension point 653 | $(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html' 654 | $(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" 655 | $(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml' 656 | $(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml" 657 | $(HIDE) rmdir "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/" || true 658 | .PHONY: uninstall-doc 659 | 660 | # Cleaning #################################################################### 661 | # 662 | # There rules can be extended in Makefile.local 663 | # Extensions can't assume when they run. 664 | 665 | clean:: 666 | @# Extension point 667 | $(SHOW)'CLEAN' 668 | $(HIDE)rm -f $(CMOFILES) 669 | $(HIDE)rm -f $(CMIFILES) 670 | $(HIDE)rm -f $(CMAFILES) 671 | $(HIDE)rm -f $(CMXFILES) 672 | $(HIDE)rm -f $(CMXAFILES) 673 | $(HIDE)rm -f $(CMXSFILES) 674 | $(HIDE)rm -f $(OFILES) 675 | $(HIDE)rm -f $(CMXAFILES:.cmxa=.a) 676 | $(HIDE)rm -f $(MLGFILES:.mlg=.ml) 677 | $(HIDE)rm -f $(CMXFILES:.cmx=.cmt) 678 | $(HIDE)rm -f $(MLIFILES:.mli=.cmti) 679 | $(HIDE)rm -f $(ALLDFILES) 680 | $(HIDE)rm -f $(NATIVEFILES) 681 | $(HIDE)find . -name .coq-native -type d -empty -delete 682 | $(HIDE)rm -f $(VOFILES) 683 | $(HIDE)rm -f $(VOFILES:.vo=.vos) 684 | $(HIDE)rm -f $(VOFILES:.vo=.vok) 685 | $(HIDE)rm -f $(VOFILES:.vo=.v.prof.json) 686 | $(HIDE)rm -f $(VOFILES:.vo=.v.prof.json.gz) 687 | $(HIDE)rm -f $(BEAUTYFILES) $(VFILES:=.old) 688 | $(HIDE)rm -f all.ps all-gal.ps all.pdf all-gal.pdf all.glob all-mli.tex 689 | $(HIDE)rm -f $(VFILES:.v=.glob) 690 | $(HIDE)rm -f $(VFILES:.v=.tex) 691 | $(HIDE)rm -f $(VFILES:.v=.g.tex) 692 | $(HIDE)rm -f pretty-timed-success.ok 693 | $(HIDE)rm -f META 694 | $(HIDE)rm -rf html mlihtml 695 | .PHONY: clean 696 | 697 | cleanall:: clean 698 | @# Extension point 699 | $(SHOW)'CLEAN *.aux *.timing' 700 | $(HIDE)rm -f $(foreach f,$(VFILES:.v=),$(dir $(f)).$(notdir $(f)).aux) 701 | $(HIDE)rm -f $(TIME_OF_BUILD_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) 702 | $(HIDE)rm -f $(VOFILES:.vo=.v.timing) 703 | $(HIDE)rm -f $(VOFILES:.vo=.v.before-timing) 704 | $(HIDE)rm -f $(VOFILES:.vo=.v.after-timing) 705 | $(HIDE)rm -f $(VOFILES:.vo=.v.timing.diff) 706 | $(HIDE)rm -f .lia.cache .nia.cache 707 | .PHONY: cleanall 708 | 709 | archclean:: 710 | @# Extension point 711 | $(SHOW)'CLEAN *.cmx *.o' 712 | $(HIDE)rm -f $(NATIVEFILES) 713 | $(HIDE)rm -f $(CMOFILES:%.cmo=%.cmx) 714 | .PHONY: archclean 715 | 716 | 717 | # Compilation rules ########################################################### 718 | 719 | $(MLIFILES:.mli=.cmi): %.cmi: %.mli 720 | $(SHOW)'CAMLC -c $<' 721 | $(HIDE)$(TIMER) $(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) $< 722 | 723 | $(MLGFILES:.mlg=.ml): %.ml: %.mlg 724 | $(SHOW)'COQPP $<' 725 | $(HIDE)$(COQPP) $< 726 | 727 | # Stupid hack around a deficient syntax: we cannot concatenate two expansions 728 | $(filter %.cmo, $(MLFILES:.ml=.cmo) $(MLGFILES:.mlg=.cmo)): %.cmo: %.ml 729 | $(SHOW)'CAMLC -c $<' 730 | $(HIDE)$(TIMER) $(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) $< 731 | 732 | # Same hack 733 | $(filter %.cmx, $(MLFILES:.ml=.cmx) $(MLGFILES:.mlg=.cmx)): %.cmx: %.ml 734 | $(SHOW)'CAMLOPT -c $(FOR_PACK) $<' 735 | $(HIDE)$(TIMER) $(CAMLOPTC) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) $(FOR_PACK) $< 736 | 737 | 738 | $(MLLIBFILES:.mllib=.cmxs): %.cmxs: %.cmxa 739 | $(SHOW)'CAMLOPT -shared -o $@' 740 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) \ 741 | -shared -o $@ $< 742 | 743 | $(MLLIBFILES:.mllib=.cma): %.cma: | %.mllib 744 | $(SHOW)'CAMLC -a -o $@' 745 | $(HIDE)$(TIMER) $(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) -a -o $@ $^ 746 | 747 | $(MLLIBFILES:.mllib=.cmxa): %.cmxa: | %.mllib 748 | $(SHOW)'CAMLOPT -a -o $@' 749 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) -a -o $@ $^ 750 | 751 | 752 | $(MLPACKFILES:.mlpack=.cmxs): %.cmxs: %.cmxa 753 | $(SHOW)'CAMLOPT -shared -o $@' 754 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) \ 755 | -shared -o $@ $< 756 | 757 | $(MLPACKFILES:.mlpack=.cmxa): %.cmxa: %.cmx | %.mlpack 758 | $(SHOW)'CAMLOPT -a -o $@' 759 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) -a -o $@ $< 760 | 761 | $(MLPACKFILES:.mlpack=.cma): %.cma: %.cmo | %.mlpack 762 | $(SHOW)'CAMLC -a -o $@' 763 | $(HIDE)$(TIMER) $(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) -a -o $@ $^ 764 | 765 | $(MLPACKFILES:.mlpack=.cmo): %.cmo: | %.mlpack 766 | $(SHOW)'CAMLC -pack -o $@' 767 | $(HIDE)$(TIMER) $(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) -pack -o $@ $^ 768 | 769 | $(MLPACKFILES:.mlpack=.cmx): %.cmx: | %.mlpack 770 | $(SHOW)'CAMLOPT -pack -o $@' 771 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) -pack -o $@ $^ 772 | 773 | # This rule is for _CoqProject with no .mllib nor .mlpack 774 | $(filter-out $(MLLIBFILES:.mllib=.cmxs) $(MLPACKFILES:.mlpack=.cmxs) $(addsuffix .cmxs,$(PACKEDFILES)) $(addsuffix .cmxs,$(LIBEDFILES)),$(MLFILES:.ml=.cmxs) $(MLGFILES:.mlg=.cmxs)): %.cmxs: %.cmx 775 | $(SHOW)'[deprecated,use-mllib-or-mlpack] CAMLOPT -shared -o $@' 776 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(FINDLIBPKGS) \ 777 | -shared -o $@ $< 778 | 779 | # can't make 780 | # https://www.gnu.org/software/make/manual/make.html#Static-Pattern 781 | # work with multiple target rules 782 | # so use eval in a loop instead 783 | # with grouped targets https://www.gnu.org/software/make/manual/make.html#Multiple-Targets 784 | # if available (GNU Make >= 4.3) 785 | ifneq (,$(filter grouped-target,$(.FEATURES))) 786 | define globvorule= 787 | 788 | # take care to $$ variables using $< etc 789 | $(1).vo $(1).glob &: $(1).v | $$(VDFILE) 790 | $$(SHOW)ROCQ compile $(1).v 791 | $$(HIDE)$$(TIMER) $$(ROCQ) compile $$(COQDEBUG) $$(TIMING_ARG) $$(PROFILE_ARG) $$(COQFLAGS) $$(COQLIBS) $(1).v 792 | $$(HIDE)$$(PROFILE_ZIP) 793 | ifeq ($(COQDONATIVE), "yes") 794 | $$(SHOW)COQNATIVE $(1).vo 795 | $$(HIDE)$$(call TIMER,$(1).vo.native) $$(COQNATIVE) $$(COQLIBS) $(1).vo 796 | endif 797 | 798 | endef 799 | else 800 | 801 | $(VOFILES): %.vo: %.v | $(VDFILE) 802 | $(SHOW)ROCQ compile $< 803 | $(HIDE)$(TIMER) $(ROCQ) compile $(COQDEBUG) $(TIMING_ARG) $(PROFILE_ARG) $(COQFLAGS) $(COQLIBS) $< 804 | $(HIDE)$(PROFILE_ZIP) 805 | ifeq ($(COQDONATIVE), "yes") 806 | $(SHOW)COQNATIVE $@ 807 | $(HIDE)$(call TIMER,$@.native) $(COQNATIVE) $(COQLIBS) $@ 808 | endif 809 | 810 | # this is broken :( todo fix if we ever find a solution that doesn't need grouped targets 811 | $(GLOBFILES): %.glob: %.v 812 | $(SHOW)'ROCQ compile $< (for .glob)' 813 | $(HIDE)$(TIMER) $(ROCQ) compile $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 814 | 815 | endif 816 | 817 | $(foreach vfile,$(VFILES:.v=),$(eval $(call globvorule,$(vfile)))) 818 | 819 | $(VFILES:.v=.vos): %.vos: %.v 820 | $(SHOW)ROCQ compile -vos $< 821 | $(HIDE)$(TIMER) $(ROCQ) compile -vos $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 822 | 823 | $(VFILES:.v=.vok): %.vok: %.v 824 | $(SHOW)ROCQ compile -vok $< 825 | $(HIDE)$(TIMER) $(ROCQ) compile -vok $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 826 | 827 | $(addsuffix .timing.diff,$(VFILES)): %.timing.diff : %.before-timing %.after-timing 828 | $(SHOW)PYTHON TIMING-DIFF $*.{before,after}-timing 829 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-single-time-diff BEFORE=$*.before-timing AFTER=$*.after-timing TIME_OF_PRETTY_BUILD_FILE="$@" 830 | 831 | $(BEAUTYFILES): %.v.beautified: %.v 832 | $(SHOW)'BEAUTIFY $<' 833 | $(HIDE)$(TIMER) $(ROCQ) compile $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -beautify $< 834 | 835 | $(TEXFILES): %.tex: %.v 836 | $(SHOW)'COQDOC -latex $<' 837 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex $< -o $@ 838 | 839 | $(GTEXFILES): %.g.tex: %.v 840 | $(SHOW)'COQDOC -latex -g $<' 841 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex -g $< -o $@ 842 | 843 | $(HTMLFILES): %.html: %.v %.glob 844 | $(SHOW)'COQDOC -html $<' 845 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -html $< -o $@ 846 | 847 | $(GHTMLFILES): %.g.html: %.v %.glob 848 | $(SHOW)'COQDOC -html -g $<' 849 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -html -g $< -o $@ 850 | 851 | # Dependency files ############################################################ 852 | 853 | ifndef MAKECMDGOALS 854 | -include $(ALLDFILES) 855 | else 856 | ifneq ($(filter-out archclean clean cleanall printenv make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff,$(MAKECMDGOALS)),) 857 | -include $(ALLDFILES) 858 | endif 859 | endif 860 | 861 | .SECONDARY: $(ALLDFILES) 862 | 863 | redir_if_ok = > "$@" || ( RV=$$?; rm -f "$@"; exit $$RV ) 864 | 865 | GENMLFILES:=$(MLGFILES:.mlg=.ml) 866 | $(addsuffix .d,$(ALLSRCFILES)): $(GENMLFILES) 867 | 868 | $(addsuffix .d,$(MLIFILES)): %.mli.d: %.mli 869 | $(SHOW)'CAMLDEP $<' 870 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 871 | 872 | $(addsuffix .d,$(MLGFILES)): %.mlg.d: %.ml 873 | $(SHOW)'CAMLDEP $<' 874 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 875 | 876 | $(addsuffix .d,$(MLFILES)): %.ml.d: %.ml 877 | $(SHOW)'CAMLDEP $<' 878 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 879 | 880 | $(addsuffix .d,$(MLLIBFILES)): %.mllib.d: %.mllib 881 | $(SHOW)'OCAMLLIBDEP $<' 882 | $(HIDE)$(OCAMLLIBDEP) -c $(OCAMLLIBS) "$<" $(redir_if_ok) 883 | 884 | $(addsuffix .d,$(MLPACKFILES)): %.mlpack.d: %.mlpack 885 | $(SHOW)'OCAMLLIBDEP $<' 886 | $(HIDE)$(OCAMLLIBDEP) -c $(OCAMLLIBS) "$<" $(redir_if_ok) 887 | 888 | # If this makefile is created using a _CoqProject we have coqdep get 889 | # options from it. This avoids argument length limits for pathological 890 | # projects. Note that extra options might be on the command line. 891 | VDFILE_FLAGS:=$(if _RocqProject,-f _RocqProject,) $(CMDLINE_COQLIBS) $(CMDLINE_VFILES) 892 | 893 | $(VDFILE): _RocqProject $(VFILES) 894 | $(SHOW)'ROCQ DEP VFILES' 895 | $(HIDE)$(TIMER) $(COQDEP) -vos -dyndep var $(VDFILE_FLAGS) $(redir_if_ok) 896 | 897 | # Misc ######################################################################## 898 | 899 | byte: 900 | $(HIDE)$(MAKE) all "OPT:=-byte" -f "$(SELF)" 901 | .PHONY: byte 902 | 903 | opt: 904 | $(HIDE)$(MAKE) all "OPT:=-opt" -f "$(SELF)" 905 | .PHONY: opt 906 | 907 | # This is deprecated. To extend this makefile use 908 | # extension points and Makefile.local 909 | printenv:: 910 | $(warning printenv is deprecated) 911 | $(warning write extensions in Makefile.local or include Makefile.conf) 912 | @echo 'COQLIB = $(COQLIB)' 913 | @echo 'COQCORELIB = $(COQCORELIB)' 914 | @echo 'DOCDIR = $(DOCDIR)' 915 | @echo 'OCAMLFIND = $(OCAMLFIND)' 916 | @echo 'HASNATDYNLINK = $(HASNATDYNLINK)' 917 | @echo 'SRC_SUBDIRS = $(SRC_SUBDIRS)' 918 | @echo 'COQ_SRC_SUBDIRS = $(COQ_SRC_SUBDIRS)' 919 | @echo 'COQCORE_SRC_SUBDIRS = $(COQCORE_SRC_SUBDIRS)' 920 | @echo 'OCAMLFIND = $(OCAMLFIND)' 921 | @echo 'PP = $(PP)' 922 | @echo 'COQFLAGS = $(COQFLAGS)' 923 | @echo 'COQLIB = $(COQLIBS)' 924 | @echo 'COQLIBINSTALL = $(COQLIBINSTALL)' 925 | @echo 'COQDOCINSTALL = $(COQDOCINSTALL)' 926 | .PHONY: printenv 927 | 928 | # Generate a .merlin file. If you need to append directives to this 929 | # file you can extend the merlin-hook target in Makefile.local 930 | .merlin: 931 | $(SHOW)'FILL .merlin' 932 | $(HIDE)echo 'FLG $(COQMF_CAMLFLAGS)' > .merlin 933 | $(HIDE)echo 'B $(COQCORELIB)' >> .merlin 934 | $(HIDE)echo 'S $(COQCORELIB)' >> .merlin 935 | $(HIDE)$(foreach d,$(COQCORE_SRC_SUBDIRS), \ 936 | echo 'B $(COQCORELIB)$(d)' >> .merlin;) 937 | $(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \ 938 | echo 'S $(COQLIB)$(d)' >> .merlin;) 939 | $(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'B $(d)' >> .merlin;) 940 | $(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'S $(d)' >> .merlin;) 941 | $(HIDE)$(MAKE) merlin-hook -f "$(SELF)" 942 | .PHONY: merlin 943 | 944 | merlin-hook:: 945 | @# Extension point 946 | .PHONY: merlin-hook 947 | 948 | # prints all variables 949 | debug: 950 | $(foreach v,\ 951 | $(sort $(filter-out $(INITIAL_VARS) INITIAL_VARS,\ 952 | $(.VARIABLES))),\ 953 | $(info $(v) = $($(v)))) 954 | .PHONY: debug 955 | 956 | .DEFAULT_GOAL := all 957 | 958 | # Users can create Makefile.local-late to hook into double-colon rules 959 | # or add other needed Makefile code, using defined 960 | # variables if necessary. 961 | -include Makefile.local-late 962 | 963 | # Local Variables: 964 | # mode: makefile-gmake 965 | # End: 966 | -------------------------------------------------------------------------------- /PermutationHelpers.v: -------------------------------------------------------------------------------- 1 | (** From the Coq manual, updated to work with the current (8.4) 2 | standard library *) 3 | From Stdlib Require Import List. 4 | From Stdlib Require Import Permutation. 5 | Ltac permutAux n := 6 | match goal with 7 | | |- (Permutation ?l ?l) => reflexivity 8 | | |- (Permutation (?a :: ?l1) (?a :: ?l2)) => 9 | let newn := eval compute in (length l1) in 10 | (apply perm_skip; permutAux newn) 11 | | |- (Permutation (?a :: ?l1) ?l2) => 12 | match eval compute in n with 13 | | 1 => fail 14 | | _ => 15 | let l1' := constr:(l1 ++ a :: nil) in 16 | (apply (@perm_trans _ (a :: l1) l1' l2); 17 | [ apply Permutation_cons_append | compute; permutAux (pred n) ]) 18 | end 19 | end. 20 | 21 | (** Permutation solver *) 22 | Ltac permut := 23 | match goal with 24 | | |- (Permutation ?l1 ?l2) => 25 | match eval compute in (length l1 = length l2) with 26 | | (?n = ?n) => permutAux n 27 | end 28 | end. 29 | 30 | (* 31 | Notation "[ x , .. , y ]" := (x :: .. (y :: nil) ..) : list_scope. 32 | Lemma fu : Permutation [1,2,3] [3,1,2]. 33 | Proof. permut. Qed. 34 | *) -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | A simple Coq development of linear logic. The library is inspired by 2 | the paper, J. Power and C. Webster, **"Working with linear logic in 3 | coq,"** in _The 12th International Conference on Theorem Proving in 4 | Higher Order Logics_, 1999. 5 | 6 | This particular devlopment is based more immediately on that used in 7 | A. Cowley and C.J. Taylor, **"Stream-oriented robotics programming: 8 | The design of roshask,"** in _Proceedings of the IEEE/RJS 9 | International Conference on Intelligent Robots and Systems (IROS)_, 10 | 2011. [pdf](http://www.seas.upenn.edu/~acowley/papers/TowardsLinear.pdf) 11 | 12 | Tested with [the Rocq Prover](https://rocq-prover.org/) 9.0. 13 | 14 | An example command to build the documentation, 15 | 16 | rocq doc --no-lib-name -d doc *.v --utf8 --no-index 17 | 18 | (this requires that the directory `doc` exists). 19 | -------------------------------------------------------------------------------- /Sig.v: -------------------------------------------------------------------------------- 1 | (** A simple module type for modules parameterized by a single 2 | type. *) 3 | Module Type Sig. 4 | Parameter A : Type. 5 | End Sig. 6 | -------------------------------------------------------------------------------- /_RocqProject: -------------------------------------------------------------------------------- 1 | -R . LinearLogic 2 | Sig.v 3 | LinProp.v 4 | LinLog.v 5 | PermutationHelpers.v 6 | Context.v 7 | BlocksWorldExample1.v 8 | 9 | --------------------------------------------------------------------------------