| author | wenzelm | 
| Mon, 06 Feb 2023 11:05:35 +0100 | |
| changeset 77205 | a197d583bf9f | 
| parent 76670 | b04d45bebbc5 | 
| child 78728 | 72631efa3821 | 
| permissions | -rw-r--r-- | 
| 62887 | 1 | (* Title: Pure/ML_Bootstrap.thy | 
| 62868 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 2 | Author: Makarius | 
| 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 3 | |
| 62887 | 4 | ML bootstrap environment -- with access to low-level structures! | 
| 62868 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 5 | *) | 
| 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 6 | |
| 62887 | 7 | theory ML_Bootstrap | 
| 62880 | 8 | imports Pure | 
| 62868 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 9 | begin | 
| 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 10 | |
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 11 | ML \<open> | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 12 | #allStruct ML_Name_Space.global () |> List.app (fn (name, _) => | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 13 | if member (op =) ML_Name_Space.hidden_structures name then | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 14 |       ML (Input.string ("structure " ^ name ^ " = " ^ name))
 | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 15 | else ()); | 
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62930diff
changeset | 16 | structure Output_Primitives = Output_Primitives_Virtual; | 
| 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62930diff
changeset | 17 | structure Thread_Data = Thread_Data_Virtual; | 
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 18 | structure PolyML = PolyML; | 
| 68918 
3a0db30e5d87
simplified signature (again, see 751bcf0473a7): e.g. relevant for non-Isabelle ML environments;
 wenzelm parents: 
68819diff
changeset | 19 | fun ML_system_pp (_: FixedInt.int -> 'a -> 'b -> PolyML.pretty) = (); | 
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62930diff
changeset | 20 | |
| 68819 | 21 | Proofterm.proofs := 0; | 
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62930diff
changeset | 22 | |
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 23 | Context.setmp_generic_context NONE | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 24 | ML \<open> | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 25 | List.app ML_Name_Space.forget_structure ML_Name_Space.hidden_structures; | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 26 | structure PolyML = | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 27 | struct | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 28 | val pointerEq = pointer_eq; | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 29 | structure IntInf = PolyML.IntInf; | 
| 68918 
3a0db30e5d87
simplified signature (again, see 751bcf0473a7): e.g. relevant for non-Isabelle ML environments;
 wenzelm parents: 
68819diff
changeset | 30 | datatype context = datatype PolyML.context; | 
| 
3a0db30e5d87
simplified signature (again, see 751bcf0473a7): e.g. relevant for non-Isabelle ML environments;
 wenzelm parents: 
68819diff
changeset | 31 | datatype pretty = datatype PolyML.pretty; | 
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 32 | end; | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 33 | \<close> | 
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62930diff
changeset | 34 | \<close> | 
| 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62930diff
changeset | 35 | |
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 36 | setup \<open>Context.theory_map ML_Env.bootstrap_name_space\<close> | 
| 68819 | 37 | |
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68812diff
changeset | 38 | declare [[ML_read_global = false]] | 
| 62868 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 39 | |
| 
61a691db1c4d
support for ML project ROOT file, with imitation of ML "use" commands;
 wenzelm parents: diff
changeset | 40 | end |