Skip to content

Commit 72b2b70

Browse files
authored
Merge pull request #115 from aqjune-aws/ocaml5
Add update_database for OCaml 5, fix a bug in search, add make switch-5
2 parents a58b904 + aea86e7 commit 72b2b70

8 files changed

+190
-49
lines changed

Makefile

+13-5
Original file line numberDiff line numberDiff line change
@@ -59,12 +59,20 @@ switch:; \
5959
opam install -y zarith ledit ; \
6060
opam pin -y add camlp5 8.03.00
6161

62+
switch-5:; \
63+
opam update ; \
64+
opam switch create . ocaml-base-compiler.5.2.0 ; \
65+
eval $(opam env) ; \
66+
opam install -y zarith ledit ; \
67+
opam pin -y add camlp5 8.03.00
68+
6269
# Choose an appropriate "update_database.ml" file
6370

64-
update_database.ml:; if [ ${OCAML_VERSION} = "4.14" ] ; \
65-
then cp update_database_4.14.ml update_database.ml ; \
66-
else cp update_database_${OCAML_UNARY_VERSION}.ml update_database.ml ; \
67-
fi
71+
update_database.ml:; \
72+
if [ ${OCAML_VERSION} = "4.14" ] ; \
73+
then cp update_database/update_database_4.14.ml update_database.ml ; \
74+
else cp update_database/update_database_${OCAML_UNARY_VERSION}.ml update_database.ml ; \
75+
fi
6876

6977
# Build the camlp4 syntax extension file (camlp5 for OCaml >= 3.10)
7078

@@ -174,7 +182,7 @@ unit_tests.native: unit_tests_inlined.ml hol_lib.cmx inline_load.ml hol.sh ; \
174182
ocamlfind ocamlopt -package zarith -linkpkg -pp "`./hol.sh -pp`" \
175183
-I . bignum.cmx hol_loader.cmx hol_lib.cmx unit_tests_inlined.ml -o unit_tests.native
176184

177-
default: hol_lib.cma unit_tests.byte unit_tests.native
185+
default: hol_lib.cma hol_lib.cmxa unit_tests.byte unit_tests.native
178186
endif
179187

180188
# TODO: update this and hol.* commands to use one of checkpointing tools

Mizarlight/make.ml

+1-1
Original file line numberDiff line numberDiff line change
@@ -21,7 +21,7 @@ let e tac =
2121

2222
Topdirs.dir_directory (!hol_dir);;
2323

24-
Topdirs.load_file Format.std_formatter
24+
Toploop.load_file Format.std_formatter
2525
(Filename.concat (!hol_dir) "Mizarlight/pa_f.cmo");;
2626

2727
List.iter (fun s -> Hashtbl.add (Pa_j.ht) s true)

README

+11-14
Original file line numberDiff line numberDiff line change
@@ -27,7 +27,7 @@ Refer to the reference manual for more details of individual functions:
2727
The Objective CAML (OCaml) implementation is a prerequisite for running
2828
HOL Light. HOL Light should work with any recent version of OCaml; I've
2929
tried it on at least 3.04, 3.06, 3.07+2, 3.08.1, 3.09.3, 3.10.0, 3.11.2,
30-
4.00, 4.05 and 4.14.
30+
4.00, 4.05 and 4.14 and 5.2.0.
3131

3232
1. OCaml: there are packages for many Linux distributions. For
3333
example, on a debian derivative like Ubuntu, you may just need
@@ -142,21 +142,18 @@ have installed. As of 2024, there are three programs you can use.
142142
HOL Light does not have convenient commands or scripts to exploit DMTCP,
143143
but you can proceed as follows:
144144

145-
1. Start ocaml running under the DMTCP coordinator:
145+
1. Start hol.sh running under the DMTCP coordinator and wait until
146+
HOL Light is loaded:
146147

147-
dmtcp_launch ocaml
148+
dmtcp_launch ./hol.sh
148149

149-
2. Use ocaml to load HOL Light as usual, for example:
150-
151-
#use "hol.ml";;
152-
153-
3. From another terminal, issue the checkpoint command:
150+
2. From another terminal, issue the checkpoint command:
154151

155152
dmtcp_command -kc
156153

157-
This will kill the ocaml process once checkpointing is done.
154+
This will kill the process once checkpointing is done.
158155

159-
4. Step 3 created a checkpoint of the OCaml process and
156+
3. Step 2 created a checkpoint of the OCaml process and
160157
a shell script to invoke it, both in the directory in
161158
which ocaml was started. Running that should restore
162159
the OCaml process with all your state and bindings:
@@ -206,7 +203,7 @@ checkpointing programs.
206203
DEPENDENCIES
207204

208205
1. zarith or num: The HOL Light system uses the OCaml "Num" library
209-
or "Zarith" library for rational arithmetic. If OCaml 4.14 is used,
206+
or "Zarith" library for rational arithmetic. If OCaml 4.14 or above is used,
210207
HOL Light will use Zarith. You can install it using the OCaml package
211208
manager "opam" by
212209

@@ -234,8 +231,8 @@ checkpointing programs.
234231
2. camlp5: this is needed to run HOL Light under any OCaml >= 3.10.
235232
Somtimes you need a recent version of camlp5 to be compatible with
236233
your OCaml. For example, OCaml 4.05 is compatible with camlp5 7.10 and
237-
OCaml 4.14 is compatible with camlp5 8.02 and 8.03. I recommend downloading
238-
the sources for a recent version from
234+
OCaml 4.14 and above is compatible with camlp5 8.02 and 8.03. I recommend
235+
downloading the sources for a recent version from
239236

240237
https://github.com/camlp5/camlp5/releases ('tags' tab has full series)
241238

@@ -269,7 +266,7 @@ HOL Light will only work on OCaml 4.14 or above.
269266

270267
To compile an OCaml file that opens Hol_lib using OCaml bytecode compiler,
271268
use the following command. For OCaml native compiler, replace ocamlc with
272-
ocamlopt.
269+
ocamlopt and .cmo with .cmx.
273270

274271
ocamlfind ocamlc -package zarith -linkpkg -pp "`./hol.sh -pp`" \
275272
-I (HOL dir) (HOL dir)/bignum.cmo (HOL dir)/hol_loader.cmo \
File renamed without changes.

update_database_4.14.ml update_database/update_database_4.14.ml

+13-8
Original file line numberDiff line numberDiff line change
@@ -26,14 +26,16 @@ let rec get_simple_type = function
2626

2727
(* Execute any OCaml expression given as a string. *)
2828

29-
let exec = ignore o Toploop.execute_phrase false Format.std_formatter
30-
o !Toploop.parse_toplevel_phrase o Lexing.from_string
29+
let exec s = (ignore o Toploop.execute_phrase false Format.std_formatter
30+
o !Toploop.parse_toplevel_phrase o Lexing.from_string) s
3131

3232
(* Evaluate any OCaml expression given as a string. *)
3333

3434
let eval n =
35-
exec ("let buf__ = ( " ^ n ^ " );;");
36-
Obj.magic (Toploop.getvalue "buf__")
35+
if String.contains n '.' then begin
36+
exec ("let buf__ = ( " ^ n ^ " );;");
37+
Obj.magic (Toploop.getvalue "buf__")
38+
end else Obj.magic (Toploop.getvalue n)
3739

3840
(* Register all theorems added since the last update. *)
3941
end
@@ -75,10 +77,13 @@ let string_of_longident lid =
7577
String.concat "." (Longident.flatten lid)
7678

7779
let all_theorems () =
78-
enum1 None []
79-
|> List.map (fun lid ->
80-
let s = string_of_longident lid in
81-
(s, (Ocaml_typing.eval s : thm)))
80+
let _ = Ocaml_typing.exec ("unset_jrh_lexer;;") in
81+
let res = enum1 None []
82+
|> List.map (fun lid ->
83+
let s = string_of_longident lid in
84+
(s, (Ocaml_typing.eval s : thm))) in
85+
let _ = Ocaml_typing.exec ("set_jrh_lexer;;") in
86+
res
8287
end
8388

8489

File renamed without changes.

update_database/update_database_5.ml

+152
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,152 @@
1+
(* ========================================================================= *)
2+
(* Create search database from OCaml / modify search database dynamically. *)
3+
(* *)
4+
(* This file assigns to "theorems", which is a list of name-theorem pairs. *)
5+
(* The core system already has such a database set up. Use this file if you *)
6+
(* want to update the database beyond the core, so you can search it. *)
7+
(* *)
8+
(* The trickery to get at the OCaml environment is due to Roland Zumkeller *)
9+
(* and Michael Farber. It works by copying some internal data structures and *)
10+
(* casting into the copy using Obj.magic. *)
11+
(* ========================================================================= *)
12+
13+
module Ocaml_typing = struct
14+
15+
open Types
16+
17+
(* If the given type is simple return its name, otherwise None. *)
18+
19+
let rec get_simple_type = function
20+
| Tlink texpr ->
21+
(match get_desc texpr with
22+
| Tconstr (Pident p,[],_) -> Some (Ident.name p)
23+
| d -> get_simple_type d)
24+
| Tconstr (Path.Pident p, [], _) -> Some (Ident.name p)
25+
| _ -> None;;
26+
27+
(* Execute any OCaml expression given as a string. *)
28+
29+
let exec s = (ignore o Toploop.execute_phrase false Format.std_formatter
30+
o !Toploop.parse_toplevel_phrase o Lexing.from_string) s
31+
32+
(* Evaluate any OCaml expression given as a string. *)
33+
34+
let eval n =
35+
if String.contains n '.' then begin
36+
exec ("let buf__ = ( " ^ n ^ " );;");
37+
Obj.magic (Toploop.getvalue "buf__")
38+
end else Obj.magic (Toploop.getvalue n)
39+
40+
(* Register all theorems added since the last update. *)
41+
end
42+
43+
module Lookuptheorems = struct
44+
open Types
45+
46+
let lid_cons lidopt id =
47+
match lidopt with
48+
None -> Longident.Lident id
49+
| Some li -> Longident.Ldot(li, id)
50+
51+
let it_val_1 lidopt s p vd acc =
52+
if (Some "thm") = Ocaml_typing.get_simple_type (get_desc vd.Types.val_type) then
53+
(lid_cons lidopt s)::acc else acc
54+
55+
let it_mod_1 lidopt s p md acc = (lid_cons lidopt s)::acc
56+
57+
let enum0 lidopt =
58+
try
59+
let vl = Env.fold_values (it_val_1 lidopt) lidopt !Toploop.toplevel_env [] in
60+
let ml = Env.fold_modules (it_mod_1 lidopt) lidopt !Toploop.toplevel_env [] in
61+
(vl, ml)
62+
with Not_found ->
63+
(* Looking for (Longident.Lident "Stream") raises Not_found.
64+
Stream is a deprecated alias module of "Stdlib.Stream", and the camlp-streams
65+
package that is used by pa_hol_syntax redefines Stream, which seems to
66+
confuse Env.fold_values and Env.fold_modules. *)
67+
([], [])
68+
69+
let rec enum1 lidopt acc =
70+
match enum0 lidopt with
71+
(vl, []) -> vl@acc
72+
| (vl, ml) ->
73+
List.fold_left (fun acc mlid ->
74+
enum1 (Some mlid) acc) (vl@acc) ml
75+
76+
let string_of_longident lid =
77+
String.concat "." (Longident.flatten lid)
78+
79+
let all_theorems () =
80+
let _ = Ocaml_typing.exec ("unset_jrh_lexer;;") in
81+
let res = enum1 None []
82+
|> List.map (fun lid ->
83+
let s = string_of_longident lid in
84+
(s, (Ocaml_typing.eval s : thm))) in
85+
let _ = Ocaml_typing.exec ("set_jrh_lexer;;") in
86+
res
87+
end
88+
89+
90+
let update_database () =
91+
theorems := Lookuptheorems.all_theorems()
92+
93+
(* ------------------------------------------------------------------------- *)
94+
(* Put an assignment of a theorem database in the named file. *)
95+
(* ------------------------------------------------------------------------- *)
96+
97+
let make_database_assignment filename =
98+
update_database();
99+
(let allnames = uniq(sort (<) (map fst (!theorems))) in
100+
let names = subtract allnames ["it"] in
101+
let entries = map (fun n -> "\""^n^"\","^n) names in
102+
let text = "needs \"help.ml\";;\n\n"^
103+
"theorems :=\n[\n"^
104+
end_itlist (fun a b -> a^";\n"^b) entries^"\n];;\n" in
105+
file_of_string filename text);;
106+
107+
(* ------------------------------------------------------------------------- *)
108+
(* Search (automatically updates) *)
109+
(* ------------------------------------------------------------------------- *)
110+
111+
let search =
112+
let rec immediatesublist l1 l2 =
113+
match (l1,l2) with
114+
[],_ -> true
115+
| _,[] -> false
116+
| (h1::t1,h2::t2) -> h1 = h2 && immediatesublist t1 t2 in
117+
let rec sublist l1 l2 =
118+
match (l1,l2) with
119+
[],_ -> true
120+
| _,[] -> false
121+
| (h1::t1,h2::t2) -> immediatesublist l1 l2 || sublist l1 t2 in
122+
let exists_subterm_satisfying p (n,th) = can (find_term p) (concl th)
123+
and name_contains s (n,th) = sublist (explode s) (explode n) in
124+
let rec filterpred tm =
125+
match tm with
126+
Comb(Var("<omit this pattern>",_),t) -> not o filterpred t
127+
| Comb(Var("<match theorem name>",_),Var(pat,_)) -> name_contains pat
128+
| Comb(Var("<match aconv>",_),pat) -> exists_subterm_satisfying (aconv pat)
129+
| pat -> exists_subterm_satisfying (can (term_match [] pat)) in
130+
fun pats ->
131+
update_database();
132+
let triv,nontriv = partition is_var pats in
133+
(if triv <> [] then
134+
warn true
135+
("Ignoring plain variables in search: "^
136+
end_itlist (fun s t -> s^", "^t) (map (fst o dest_var) triv))
137+
else ());
138+
(if nontriv = [] && triv <> [] then []
139+
else sort (increasing fst)
140+
(itlist (filter o filterpred) pats (!theorems)));;
141+
142+
(* ------------------------------------------------------------------------- *)
143+
(* Update to bring things back to current state. *)
144+
(* ------------------------------------------------------------------------- *)
145+
146+
update_database();;
147+
148+
(* This printf checks whether standard modules like Printf are still alive
149+
after update_database.
150+
See also: https://github.com/ocaml/ocaml/issues/12271 *)
151+
Printf.printf "update_database.ml loaded! # theorems: %d\n"
152+
(List.length !theorems);;

update_database_5.ml

-21
This file was deleted.

0 commit comments

Comments
 (0)