Skip to content

Commit

Permalink
Merge with recent trunk and new version of universe polymorphism.
Browse files Browse the repository at this point in the history
Conflicts:
	kernel/declarations.ml
  • Loading branch information
mattam82 committed Apr 3, 2013
2 parents e83990d + 9441842 commit 44d41f2
Show file tree
Hide file tree
Showing 473 changed files with 16,804 additions and 9,145 deletions.
1 change: 1 addition & 0 deletions .gitignore
Expand Up @@ -36,6 +36,7 @@
*.v.ps *.v.ps
*.v.html *.v.html
*.stamp *.stamp
*.native
revision revision
TAGS TAGS
.DS_Store .DS_Store
Expand Down
9 changes: 9 additions & 0 deletions CHANGES
Expand Up @@ -8,12 +8,15 @@ Vernacular commands
A flag "Set/Unset Record Elimination Schemes" allows to change this. A flag "Set/Unset Record Elimination Schemes" allows to change this.
The tactic "induction" on a record is now actually doing "destruct". The tactic "induction" on a record is now actually doing "destruct".
- The "Open Scope" command can now be given also a delimiter (e.g. Z). - The "Open Scope" command can now be given also a delimiter (e.g. Z).
- The "Definition" command now allows the "Local" modifier.
- The "Let" command can now define local (co)fixpoints.


Specification Language Specification Language


- The syntax "x -> y" is now declared at level 99. In particular, it has - The syntax "x -> y" is now declared at level 99. In particular, it has
now a lower priority than "<->": "A -> B <-> C" is now "A -> (B <-> C)" now a lower priority than "<->": "A -> B <-> C" is now "A -> (B <-> C)"
(possible source of incompatibilities) (possible source of incompatibilities)
- Slight changes in unification error messages.


Tactics Tactics


Expand All @@ -27,6 +30,12 @@ Tactics
- Similarly, "intuition" has been made more uniform and, where it now - Similarly, "intuition" has been made more uniform and, where it now
fails, "dintuition" can be used. (possible source of incompatibilities) fails, "dintuition" can be used. (possible source of incompatibilities)
- Tactic notations can now be defined locally to a module (use "Local" prefix). - Tactic notations can now be defined locally to a module (use "Local" prefix).
- Tactic "red" now reduces head beta-iota redexes (potential source of
rare incompatibilities).
- Tactic "hnf" now reduces inner beta-iota redexes
(potential source of rare incompatibilities).
- Tactic "intro H" now reduces beta-iota redexes if these hide a product
(potential source of rare incompatibilities).


Program Program


Expand Down
2 changes: 1 addition & 1 deletion Makefile
Expand Up @@ -222,7 +222,7 @@ cleanconfig:
distclean: clean cleanconfig distclean: clean cleanconfig


voclean: voclean:
find theories plugins test-suite -name '*.vo' -o -name '*.glob' | xargs rm -f find theories plugins test-suite -name '*.vo' -o -name '*.glob' -o -name "*.cmxs" -o -name "*.native" -o -name "*.cmx" -o -name "*.cmi" -o -name "*.o" | xargs rm -f


devdocclean: devdocclean:
find . -name '*.dep.ps' -o -name '*.dot' | xargs rm -f find . -name '*.dep.ps' -o -name '*.dot' | xargs rm -f
Expand Down
28 changes: 8 additions & 20 deletions Makefile.build
Expand Up @@ -125,7 +125,7 @@ endif


PR_O := $(if $(READABLE_ML4),pr_o.cmo,pr_dump.cmo) # works also with new camlp4 PR_O := $(if $(READABLE_ML4),pr_o.cmo,pr_dump.cmo) # works also with new camlp4


SYSMOD:=str unix SYSMOD:=str unix dynlink
SYSCMA:=$(addsuffix .cma,$(SYSMOD)) SYSCMA:=$(addsuffix .cma,$(SYSMOD))
SYSCMXA:=$(addsuffix .cmxa,$(SYSMOD)) SYSCMXA:=$(addsuffix .cmxa,$(SYSMOD))


Expand Down Expand Up @@ -370,7 +370,8 @@ install-ide-info:


.PHONY: validate check test-suite $(ALLSTDLIB).v .PHONY: validate check test-suite $(ALLSTDLIB).v


VALIDOPTS=-silent -o -m
VALIDOPTS=$(if $(VERBOSE),,-silent) -o -m


validate:: $(CHICKEN) $(ALLVO) validate:: $(CHICKEN) $(ALLVO)
$(SHOW)'COQCHK <theories & plugins>' $(SHOW)'COQCHK <theories & plugins>'
Expand Down Expand Up @@ -675,12 +676,12 @@ install-latex:
source-doc: mli-doc $(OCAMLDOCDIR)/coq.pdf source-doc: mli-doc $(OCAMLDOCDIR)/coq.pdf


$(OCAMLDOCDIR)/coq.tex:: $(DOCMLIS:.mli=.cmi) $(OCAMLDOCDIR)/coq.tex:: $(DOCMLIS:.mli=.cmi)
$(OCAMLDOC) -latex -I $(MYCAMLP4LIB) $(MLINCLUDES)\ $(OCAMLDOC) -latex -rectypes -I $(MYCAMLP4LIB) $(MLINCLUDES)\
$(DOCMLIS) -t "Coq mlis documentation" \ $(DOCMLIS) -t "Coq mlis documentation" \
-intro $(OCAMLDOCDIR)/docintro -o $@ -intro $(OCAMLDOCDIR)/docintro -o $@


mli-doc:: $(DOCMLIS:.mli=.cmi) mli-doc:: $(DOCMLIS:.mli=.cmi)
$(OCAMLDOC) -html -I $(MYCAMLP4LIB) $(MLINCLUDES)\ $(OCAMLDOC) -html -rectypes -I $(MYCAMLP4LIB) $(MLINCLUDES)\
$(DOCMLIS) -d $(OCAMLDOCDIR)/html -colorize-code \ $(DOCMLIS) -d $(OCAMLDOCDIR)/html -colorize-code \
-t "Coq mlis documentation" -intro $(OCAMLDOCDIR)/docintro \ -t "Coq mlis documentation" -intro $(OCAMLDOCDIR)/docintro \
-css-style style.css -css-style style.css
Expand All @@ -693,9 +694,9 @@ ml-dot:: $(MLEXTRAFILES)
$(DOT) -Tpng $< -o $@ $(DOT) -Tpng $< -o $@


%_types.dot: %.mli %_types.dot: %.mli
$(OCAMLDOC) $(MLINCLUDES) $(ODOCDOTOPTS) -dot-types -o $@ $< $(OCAMLDOC) -rectypes $(MLINCLUDES) $(ODOCDOTOPTS) -dot-types -o $@ $<


OCAMLDOC_MLLIBD = $(OCAMLDOC) $(MLINCLUDES) $(ODOCDOTOPTS) -o $@ \ OCAMLDOC_MLLIBD = $(OCAMLDOC) -rectypes $(MLINCLUDES) $(ODOCDOTOPTS) -o $@ \
$(foreach lib,$(|:.mllib.d=_MLLIB_DEPENDENCIES),$(addsuffix .ml,$($(lib)))) $(foreach lib,$(|:.mllib.d=_MLLIB_DEPENDENCIES),$(addsuffix .ml,$($(lib))))


%.dot: | %.mllib.d %.dot: | %.mllib.d
Expand All @@ -714,7 +715,7 @@ tactics/tactics.dot: | tactics/tactics.mllib.d tactics/hightactics.mllib.d
$(OCAMLDOC_MLLIBD) $(OCAMLDOC_MLLIBD)


%.dot: %.mli %.dot: %.mli
$(OCAMLDOC) $(MLINCLUDES) $(ODOCDOTOPTS) -o $@ $< $(OCAMLDOC) -rectypes $(MLINCLUDES) $(ODOCDOTOPTS) -o $@ $<


$(OCAMLDOCDIR)/%.pdf: $(OCAMLDOCDIR)/%.tex $(OCAMLDOCDIR)/%.pdf: $(OCAMLDOCDIR)/%.tex
(cd $(OCAMLDOCDIR) ; pdflatex $*.tex && pdflatex $*.tex) (cd $(OCAMLDOCDIR) ; pdflatex $*.tex && pdflatex $*.tex)
Expand Down Expand Up @@ -746,19 +747,6 @@ ide/coqide_main_opt.ml: ide/coqide_main.ml4 config/Makefile # no camlp4deps here
$(SHOW)'CAMLP4O $<' $(SHOW)'CAMLP4O $<'
$(HIDE)$(CAMLP4O) $(PR_O) $(CAMLP4USE) -D$(IDEOPTINT) -impl $< -o $@ $(HIDE)$(CAMLP4O) $(PR_O) $(CAMLP4USE) -D$(IDEOPTINT) -impl $< -o $@


# Specific rule for term.ml which uses -rectypes
# Since term.mli doesn't need -rectypes, this doesn't impact the other ml files

TERM:=kernel/term

$(TERM).cmo: $(TERM).ml | $(TERM).ml.d
$(SHOW)'OCAMLC -rectypes $<'
$(HIDE)$(OCAMLC) $(MLINCLUDES) $(BYTEFLAGS) -rectypes -c $<

$(TERM).cmx: $(TERM).ml | $(TERM).ml.d
$(SHOW)'OCAMLOPT -rectypes $<'
$(HIDE)$(OCAMLOPT) $(MLINCLUDES) $(OPTFLAGS) -intf-suffix .cmi -rectypes -c $<

# pretty printing of the revision number when compiling a checked out # pretty printing of the revision number when compiling a checked out
# source tree # source tree
.PHONY: revision .PHONY: revision
Expand Down
11 changes: 10 additions & 1 deletion Makefile.common
Expand Up @@ -323,9 +323,18 @@ ALLSTDLIB := test-suite/misc/universes/all_stdlib
# remove .vo, replace theories and plugins by Coq, and replace slashes by dots # remove .vo, replace theories and plugins by Coq, and replace slashes by dots
vo_to_mod = $(subst /,.,$(patsubst theories/%,Coq.%,$(patsubst plugins/%,Coq.%,$(1:.vo=)))) vo_to_mod = $(subst /,.,$(patsubst theories/%,Coq.%,$(patsubst plugins/%,Coq.%,$(1:.vo=))))


# Converting a stdlib filename into native compiler filenames
# Used for install targets
vo_to_cm = $(foreach vo,$(1),$(dir $(vo))$(subst /,_,$(patsubst theories/%,NCoq_%,$(patsubst plugins/%,NCoq_%,$(vo:.vo=.cm*)))))

vo_to_obj = $(foreach vo,$(1),$(dir $(vo))$(subst /,_,$(patsubst theories/%,NCoq_%,$(patsubst plugins/%,NCoq_%,$(vo:.vo=.o)))))

ALLMODS:=$(call vo_to_mod,$(ALLVO)) ALLMODS:=$(call vo_to_mod,$(ALLVO))


LIBFILES:=$(THEORIESVO) $(PLUGINSVO) LIBFILES:=$(THEORIESVO) $(PLUGINSVO) $(call vo_to_cm,$(THEORIESVO)) \
$(call vo_to_cm,$(PLUGINSVO)) $(call vo_to_obj,$(THEORIESVO)) \
$(call vo_to_obj,$(PLUGINSVO))

LIBFILESLIGHT:=$(THEORIESLIGHTVO) LIBFILESLIGHT:=$(THEORIESLIGHTVO)


########################################################################### ###########################################################################
Expand Down
4 changes: 0 additions & 4 deletions _tags
Expand Up @@ -22,10 +22,6 @@


<grammar/grammar.{cma,cmxa}> : use_unix <grammar/grammar.{cma,cmxa}> : use_unix


## tags for ml files

"kernel/term.ml": rectypes

## tags for camlp4 files ## tags for camlp4 files


"toplevel/whelp.ml4": use_grammar "toplevel/whelp.ml4": use_grammar
Expand Down
63 changes: 32 additions & 31 deletions checker/check.ml
Expand Up @@ -11,19 +11,19 @@ open Errors
open Util open Util
open Names open Names


let pr_dirpath dp = str (Dir_path.to_string dp) let pr_dirpath dp = str (DirPath.to_string dp)
let default_root_prefix = Dir_path.make [] let default_root_prefix = DirPath.empty
let split_dirpath d = let split_dirpath d =
let l = Dir_path.repr d in (Dir_path.make (List.tl l), List.hd l) let l = DirPath.repr d in (DirPath.make (List.tl l), List.hd l)
let extend_dirpath p id = Dir_path.make (id :: Dir_path.repr p) let extend_dirpath p id = DirPath.make (id :: DirPath.repr p)


type section_path = { type section_path = {
dirpath : string list ; dirpath : string list ;
basename : string } basename : string }
let dir_of_path p = let dir_of_path p =
Dir_path.make (List.map Id.of_string p.dirpath) DirPath.make (List.map Id.of_string p.dirpath)
let path_of_dirpath dir = let path_of_dirpath dir =
match Dir_path.repr dir with match DirPath.repr dir with
[] -> failwith "path_of_dirpath" [] -> failwith "path_of_dirpath"
| l::dir -> | l::dir ->
{dirpath=List.map Id.to_string dir;basename=Id.to_string l} {dirpath=List.map Id.to_string dir;basename=Id.to_string l}
Expand All @@ -36,7 +36,7 @@ let pr_path sp =


type library_objects type library_objects


type compilation_unit_name = Dir_path.t type compilation_unit_name = DirPath.t


type library_disk = { type library_disk = {
md_name : compilation_unit_name; md_name : compilation_unit_name;
Expand All @@ -61,10 +61,10 @@ type library_t = {


module LibraryOrdered = module LibraryOrdered =
struct struct
type t = Dir_path.t type t = DirPath.t
let compare d1 d2 = let compare d1 d2 =
Pervasives.compare Pervasives.compare
(List.rev (Dir_path.repr d1)) (List.rev (Dir_path.repr d2)) (List.rev (DirPath.repr d1)) (List.rev (DirPath.repr d2))
end end


module LibrarySet = Set.Make(LibraryOrdered) module LibrarySet = Set.Make(LibraryOrdered)
Expand All @@ -81,7 +81,7 @@ let find_library dir =
let try_find_library dir = let try_find_library dir =
try find_library dir try find_library dir
with Not_found -> with Not_found ->
error ("Unknown library " ^ (Dir_path.to_string dir)) error ("Unknown library " ^ (DirPath.to_string dir))


let library_full_filename dir = (find_library dir).library_filename let library_full_filename dir = (find_library dir).library_filename


Expand Down Expand Up @@ -113,7 +113,7 @@ let check_one_lib admit (dir,m) =
(*************************************************************************) (*************************************************************************)
(*s Load path. Mapping from physical to logical paths etc.*) (*s Load path. Mapping from physical to logical paths etc.*)


type logical_path = Dir_path.t type logical_path = DirPath.t


let load_paths = ref ([],[] : CUnix.physical_path list * logical_path list) let load_paths = ref ([],[] : CUnix.physical_path list * logical_path list)


Expand Down Expand Up @@ -153,7 +153,7 @@ let find_logical_path phys_dir =
match List.filter2 (fun p d -> p = phys_dir) physical logical with match List.filter2 (fun p d -> p = phys_dir) physical logical with
| _,[dir] -> dir | _,[dir] -> dir
| _,[] -> default_root_prefix | _,[] -> default_root_prefix
| _,l -> anomaly ("Two logical paths are associated to "^phys_dir) | _,l -> anomaly (Pp.str ("Two logical paths are associated to "^phys_dir))


let remove_load_path dir = let remove_load_path dir =
let physical, logical = !load_paths in let physical, logical = !load_paths in
Expand Down Expand Up @@ -185,7 +185,7 @@ let add_load_path (phys_path,coq_path) =
end end
| _,[] -> | _,[] ->
load_paths := (phys_path :: fst !load_paths, coq_path :: snd !load_paths) load_paths := (phys_path :: fst !load_paths, coq_path :: snd !load_paths)
| _ -> anomaly ("Two logical paths are associated to "^phys_path) | _ -> anomaly (Pp.str ("Two logical paths are associated to "^phys_path))


let load_paths_of_dir_path dir = let load_paths_of_dir_path dir =
let physical, logical = !load_paths in let physical, logical = !load_paths in
Expand Down Expand Up @@ -234,36 +234,37 @@ let locate_qualified_library qid =
(dir, file) (dir, file)
with Not_found -> raise LibNotFound with Not_found -> raise LibNotFound


let explain_locate_library_error qid = function let error_unmapped_dir qid =
| LibUnmappedDir -> let prefix = qid.dirpath in
let prefix = qid.dirpath in errorlabstrm "load_absolute_library_from"
errorlabstrm "load_absolute_library_from" (str "Cannot load " ++ pr_path qid ++ str ":" ++ spc () ++
(str "Cannot load " ++ pr_path qid ++ str ":" ++ spc () ++ str "no physical path bound to" ++ spc () ++ pr_dirlist prefix ++ fnl ())
str "no physical path bound to" ++ spc () ++ pr_dirlist prefix ++ fnl ())
| LibNotFound -> let error_lib_not_found qid =
errorlabstrm "load_absolute_library_from" errorlabstrm "load_absolute_library_from"
(str"Cannot find library " ++ pr_path qid ++ str" in loadpath") (str"Cannot find library " ++ pr_path qid ++ str" in loadpath")
| e -> raise e


let try_locate_absolute_library dir = let try_locate_absolute_library dir =
try try
locate_absolute_library dir locate_absolute_library dir
with e -> with
explain_locate_library_error (path_of_dirpath dir) e | LibUnmappedDir -> error_unmapped_dir (path_of_dirpath dir)
| LibNotFound -> error_lib_not_found (path_of_dirpath dir)


let try_locate_qualified_library qid = let try_locate_qualified_library qid =
try try
locate_qualified_library qid locate_qualified_library qid
with e -> with
explain_locate_library_error qid e | LibUnmappedDir -> error_unmapped_dir qid
| LibNotFound -> error_lib_not_found qid


(************************************************************************) (************************************************************************)
(*s Low-level interning/externing of libraries to files *) (*s Low-level interning/externing of libraries to files *)


(*s Loading from disk to cache (preparation phase) *) (*s Loading from disk to cache (preparation phase) *)


let raw_intern_library = let raw_intern_library =
snd (System.raw_extern_intern Coq_config.vo_magic_number ".vo") snd (System.raw_extern_intern Coq_config.vo_magic_number)


let with_magic_number_check f a = let with_magic_number_check f a =
try f a try f a
Expand Down Expand Up @@ -296,9 +297,9 @@ let intern_from_file (dir, f) =
let (md,table,digest) = let (md,table,digest) =
try try
let ch = with_magic_number_check raw_intern_library f in let ch = with_magic_number_check raw_intern_library f in
let (md:library_disk) = System.marshal_in ch in let (md:library_disk) = System.marshal_in f ch in
let digest = System.marshal_in ch in let digest = System.marshal_in f ch in
let table = (System.marshal_in ch : Safe_typing.LightenLibrary.table) in let table = (System.marshal_in f ch : Safe_typing.LightenLibrary.table) in
close_in ch; close_in ch;
if dir <> md.md_name then if dir <> md.md_name then
errorlabstrm "load_physical_library" errorlabstrm "load_physical_library"
Expand Down
5 changes: 4 additions & 1 deletion checker/check.mllib
@@ -1,5 +1,9 @@
Coq_config Coq_config
Int Int
Option
Store
Exninfo
Backtrace
Pp_control Pp_control
Flags Flags
Pp Pp
Expand All @@ -15,7 +19,6 @@ CList
CString CString
CArray CArray
Util Util
Option
CUnix CUnix
System System
Envars Envars
Expand Down

0 comments on commit 44d41f2

Please sign in to comment.