src/Pure/ProofGeneral/pgip_isabelle.ML
author wenzelm
Sat, 30 Dec 2006 16:08:06 +0100
changeset 21966 edab0ecfbd7c
parent 21940 fbd068dd4d29
child 22165 eaec72532dd7
permissions -rw-r--r--
removed misleading OuterLex.eq_token;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21637
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     1
(*  Title:      Pure/ProofGeneral/pgip_isabelle.ML
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     2
    ID:         $Id$
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     3
    Author:     David Aspinall
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     4
21940
fbd068dd4d29 minor tuning;
wenzelm
parents: 21637
diff changeset
     5
Prover-side PGIP abstraction: Isabelle configuration.
21637
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     6
*)
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     7
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     8
signature PGIP_ISABELLE =
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
     9
sig
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    10
    val isabelle_pgml_version_supported : string
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    11
    val isabelle_pgip_version_supported : string
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    12
    val accepted_inputs : (string * bool * (string list)) list 
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    13
end
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    14
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    15
structure PgipIsabelle : PGIP_ISABELLE =
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    16
struct
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    17
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    18
val isabelle_pgml_version_supported = "1.0";
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    19
val isabelle_pgip_version_supported = "2.0"
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    20
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    21
(** Accepted commands **)
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    22
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    23
local
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    24
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    25
    (* These element names are a subset of those in pgip_input.ML.
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    26
       They must be handled in proof_general_pgip.ML/process_pgip_element. *)
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    27
						    
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    28
    val accepted_names =
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    29
    (* not implemented: "askconfig", "forget", "restoregoal" *)
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    30
    ["askpgip","askpgml","askprefs","getpref","setpref",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    31
     "proverinit","proverexit","startquiet","stopquiet",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    32
     "pgmlsymbolson", "pgmlsymbolsoff",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    33
     "dostep", "undostep", "redostep", "abortgoal", 
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    34
     "askids", "showid", "askguise",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    35
     "parsescript",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    36
     "showproofstate", "showctxt", "searchtheorems", "setlinewidth", "viewdoc",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    37
     "doitem", "undoitem", "redoitem", "abortheory",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    38
     "retracttheory", "loadfile", "openfile", "closefile",
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    39
     "abortfile", "retractfile", "changecwd", "systemcmd"];
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    40
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    41
    fun element_async p = 
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    42
	false (* single threaded only *)
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    43
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    44
    fun supported_optional_attrs p = (case p of
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    45
					  "undostep" => ["times"]
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    46
					(* TODO: we could probably extend these too:
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    47
					| "redostep" => ["times"]
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    48
					| "undoitem" => ["times"]
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    49
					| "redoitem" => ["times"] *)
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    50
				        | _ => [])
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    51
in
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    52
val accepted_inputs = 
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    53
    (map (fn p=> (p, element_async p, supported_optional_attrs p))
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    54
         accepted_names);
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    55
end
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    56
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    57
end
a7b156c404e2 Revamped Proof General interface.
aspinall
parents:
diff changeset
    58