| author | haftmann | 
| Wed, 05 Dec 2007 14:16:14 +0100 | |
| changeset 25542 | ced4104f6c1f | 
| parent 23884 | 1d39ec4fe73f | 
| child 26005 | 431ab3907291 | 
| permissions | -rw-r--r-- | 
| 21637 | 1  | 
(* Title: Pure/ProofGeneral/pgip_isabelle.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: David Aspinall  | 
|
4  | 
||
| 22165 | 5  | 
Prover-side PGIP abstraction: Isabelle configuration and mapping to Isabelle types.  | 
| 21637 | 6  | 
*)  | 
7  | 
||
8  | 
signature PGIP_ISABELLE =  | 
|
9  | 
sig  | 
|
10  | 
val isabelle_pgml_version_supported : string  | 
|
11  | 
val isabelle_pgip_version_supported : string  | 
|
| 23603 | 12  | 
val systemid : string  | 
| 23799 | 13  | 
val accepted_inputs : (string * bool * (string list)) list  | 
| 22165 | 14  | 
|
15  | 
val location_of_position : Position.T -> PgipTypes.location  | 
|
| 22403 | 16  | 
|
17  | 
(* Additional types of objects in Isar scripts *)  | 
|
18  | 
||
19  | 
val ObjTheoryBody : PgipTypes.objtype  | 
|
20  | 
val ObjTheoryDecl : PgipTypes.objtype  | 
|
21  | 
val ObjTheoryBodySubsection : PgipTypes.objtype  | 
|
22  | 
val ObjProofBody : PgipTypes.objtype  | 
|
23  | 
val ObjFormalComment : PgipTypes.objtype  | 
|
24  | 
val ObjClass : PgipTypes.objtype  | 
|
25  | 
val ObjTheoremSet : PgipTypes.objtype  | 
|
26  | 
val ObjOracle : PgipTypes.objtype  | 
|
27  | 
val ObjLocale : PgipTypes.objtype  | 
|
| 23799 | 28  | 
|
| 21637 | 29  | 
end  | 
30  | 
||
31  | 
structure PgipIsabelle : PGIP_ISABELLE =  | 
|
32  | 
struct  | 
|
33  | 
||
| 23603 | 34  | 
val isabelle_pgml_version_supported = "2.0";  | 
| 21637 | 35  | 
val isabelle_pgip_version_supported = "2.0"  | 
| 23603 | 36  | 
val systemid = "Isabelle"  | 
| 21637 | 37  | 
|
| 23799 | 38  | 
|
| 21637 | 39  | 
(** Accepted commands **)  | 
40  | 
||
41  | 
local  | 
|
42  | 
||
43  | 
(* These element names are a subset of those in pgip_input.ML.  | 
|
44  | 
They must be handled in proof_general_pgip.ML/process_pgip_element. *)  | 
|
| 23799 | 45  | 
|
| 21637 | 46  | 
val accepted_names =  | 
47  | 
(* not implemented: "askconfig", "forget", "restoregoal" *)  | 
|
48  | 
["askpgip","askpgml","askprefs","getpref","setpref",  | 
|
49  | 
"proverinit","proverexit","startquiet","stopquiet",  | 
|
50  | 
"pgmlsymbolson", "pgmlsymbolsoff",  | 
|
| 23799 | 51  | 
"dostep", "undostep", "redostep", "abortgoal",  | 
| 21637 | 52  | 
"askids", "showid", "askguise",  | 
53  | 
"parsescript",  | 
|
54  | 
"showproofstate", "showctxt", "searchtheorems", "setlinewidth", "viewdoc",  | 
|
55  | 
"doitem", "undoitem", "redoitem", "abortheory",  | 
|
56  | 
"retracttheory", "loadfile", "openfile", "closefile",  | 
|
57  | 
"abortfile", "retractfile", "changecwd", "systemcmd"];  | 
|
58  | 
||
| 23799 | 59  | 
fun element_async p =  | 
60  | 
false (* single threaded only *)  | 
|
| 21637 | 61  | 
|
62  | 
fun supported_optional_attrs p = (case p of  | 
|
| 23799 | 63  | 
"undostep" => ["times"]  | 
64  | 
(* TODO: we could probably extend these too:  | 
|
65  | 
| "redostep" => ["times"]  | 
|
66  | 
| "undoitem" => ["times"]  | 
|
67  | 
| "redoitem" => ["times"] *)  | 
|
68  | 
| _ => [])  | 
|
| 21637 | 69  | 
in  | 
| 23799 | 70  | 
val accepted_inputs =  | 
| 21637 | 71  | 
(map (fn p=> (p, element_async p, supported_optional_attrs p))  | 
72  | 
accepted_names);  | 
|
73  | 
end  | 
|
74  | 
||
| 22165 | 75  | 
|
| 23799 | 76  | 
fun location_of_position pos =  | 
| 22165 | 77  | 
let val line = Position.line_of pos  | 
| 23799 | 78  | 
val (url,descr) =  | 
79  | 
(case Position.file_of pos of  | 
|
80  | 
NONE => (NONE, NONE)  | 
|
81  | 
| SOME fname =>  | 
|
82  | 
let val path = Path.explode fname in  | 
|
| 
23884
 
1d39ec4fe73f
simplified ThyLoad interfaces: only one additional directory;
 
wenzelm 
parents: 
23867 
diff
changeset
 | 
83  | 
case ThyLoad.check_file Path.current path of  | 
| 23799 | 84  | 
SOME _ => (SOME (PgipTypes.pgipurl_of_path path), NONE)  | 
85  | 
| NONE => (NONE, SOME fname)  | 
|
86  | 
end);  | 
|
| 23680 | 87  | 
in  | 
| 23799 | 88  | 
        { descr=descr, url=url, line=line, column=NONE, char=NONE, length=NONE }
 | 
| 23680 | 89  | 
end  | 
| 22403 | 90  | 
|
91  | 
||
92  | 
val [ObjTheoryBody,  | 
|
93  | 
ObjTheoryDecl,  | 
|
94  | 
ObjTheoryBodySubsection,  | 
|
95  | 
ObjProofBody,  | 
|
96  | 
ObjFormalComment,  | 
|
97  | 
ObjClass,  | 
|
98  | 
ObjTheoremSet,  | 
|
99  | 
ObjOracle,  | 
|
100  | 
ObjLocale] =  | 
|
| 23799 | 101  | 
map PgipTypes.ObjOther  | 
102  | 
["theory body",  | 
|
103  | 
"theory declaration",  | 
|
104  | 
"theory subsection",  | 
|
105  | 
"proof body",  | 
|
106  | 
"formal comment",  | 
|
107  | 
"class",  | 
|
108  | 
"theorem set declaration",  | 
|
109  | 
"oracle",  | 
|
110  | 
"locale"];  | 
|
| 22403 | 111  | 
|
| 21637 | 112  | 
end  |