--- a/src/Pure/Isar/isar_cmd.ML Sat Jul 28 22:01:01 2007 +0200
+++ b/src/Pure/Isar/isar_cmd.ML Sat Jul 28 22:01:06 2007 +0200
@@ -233,7 +233,7 @@
(* declarations *)
val declaration =
- ML_Context.use_let "val declaration: declaration"
+ ML_Context.use_let "val declaration: Morphism.declaration"
"Context.map_proof (LocalTheory.declaration declaration)"
#> Context.proof_map;
--- a/src/Pure/Isar/local_theory.ML Sat Jul 28 22:01:01 2007 +0200
+++ b/src/Pure/Isar/local_theory.ML Sat Jul 28 22:01:06 2007 +0200
@@ -6,7 +6,6 @@
*)
type local_theory = Proof.context;
-type declaration = morphism -> Context.generic -> Context.generic;
signature LOCAL_THEORY =
sig
--- a/src/Pure/morphism.ML Sat Jul 28 22:01:01 2007 +0200
+++ b/src/Pure/morphism.ML Sat Jul 28 22:01:06 2007 +0200
@@ -10,6 +10,7 @@
signature BASIC_MORPHISM =
sig
type morphism
+ type declaration = morphism -> Context.generic -> Context.generic
val $> : morphism * morphism -> morphism
end
@@ -51,6 +52,8 @@
term: term -> term,
fact: thm list -> thm list};
+type declaration = morphism -> Context.generic -> Context.generic;
+
fun name (Morphism {name, ...}) = name;
fun var (Morphism {var, ...}) = var;
fun typ (Morphism {typ, ...}) = typ;