discontinued obsolete alias structure ProofContext;
authorwenzelm
Wed, 12 Oct 2011 16:21:07 +0200
changeset 45128 5af3a3203a76
parent 45127 d2eb07a1e01b
child 45129 1fce03e3e8ad
discontinued obsolete alias structure ProofContext;
NEWS
src/HOL/Nitpick_Examples/minipick.ML
src/Pure/pure_setup.ML
src/Tools/Code/code_thingol.ML
--- a/NEWS	Wed Oct 12 09:16:30 2011 +0200
+++ b/NEWS	Wed Oct 12 16:21:07 2011 +0200
@@ -44,6 +44,12 @@
   zero_le_zpower_abs ~> zero_le_power_abs
 
 
+*** ML ***
+
+* Structure Proof_Context follows standard naming scheme.  Old
+ProofContext has been discontinued.  INCOMPATIBILITY.
+
+
 
 New in Isabelle2011-1 (October 2011)
 ------------------------------------
--- a/src/HOL/Nitpick_Examples/minipick.ML	Wed Oct 12 09:16:30 2011 +0200
+++ b/src/HOL/Nitpick_Examples/minipick.ML	Wed Oct 12 16:21:07 2011 +0200
@@ -381,7 +381,7 @@
 
 fun kodkod_problem_from_term ctxt total raw_card raw_infinite t =
   let
-    val thy = ProofContext.theory_of ctxt
+    val thy = Proof_Context.theory_of ctxt
     fun card (Type (@{type_name fun}, [T1, T2])) =
         reasonable_power (card T2) (card T1)
       | card (Type (@{type_name prod}, [T1, T2])) = card T1 * card T2
@@ -433,7 +433,7 @@
 
 fun minipick ctxt n t =
   let
-    val thy = ProofContext.theory_of ctxt
+    val thy = Proof_Context.theory_of ctxt
     val {total_consts, ...} = Nitpick_Isar.default_params thy []
     val totals =
       total_consts |> Option.map single |> the_default [true, false]
--- a/src/Pure/pure_setup.ML	Wed Oct 12 09:16:30 2011 +0200
+++ b/src/Pure/pure_setup.ML	Wed Oct 12 16:21:07 2011 +0200
@@ -57,6 +57,3 @@
 
 Proofterm.proofs := 0;
 
-(*legacy*)
-structure ProofContext = Proof_Context;
-
--- a/src/Tools/Code/code_thingol.ML	Wed Oct 12 09:16:30 2011 +0200
+++ b/src/Tools/Code/code_thingol.ML	Wed Oct 12 16:21:07 2011 +0200
@@ -606,7 +606,7 @@
 
 fun annotate thy algbr eqngr (c, ty) args rhs =
   let
-    val ctxt = ProofContext.init_global thy |> Config.put Type_Infer_Context.const_sorts false
+    val ctxt = Proof_Context.init_global thy |> Config.put Type_Infer_Context.const_sorts false
     val erase = map_types (fn _ => Type_Infer.anyT [])
     val reinfer = singleton (Type_Infer_Context.infer_types ctxt)
     val lhs = list_comb (Const (c, ty), map (map_types Type.strip_sorts o fst) args)