src/Pure/Isar/isar_cmd.ML
changeset 82582 bcee022fbe30
parent 81594 7e1b66416b7f
equal deleted inserted replaced
82581:0133fe6a1f55 82582:bcee022fbe30
   242 
   242 
   243 (* print theorems, terms, types etc. *)
   243 (* print theorems, terms, types etc. *)
   244 
   244 
   245 local
   245 local
   246 
   246 
   247 fun string_of_stmts ctxt args =
   247 fun pretty_stmts ctxt args =
   248   Attrib.eval_thms ctxt args
   248   Attrib.eval_thms ctxt args
   249   |> map (Element.pretty_statement ctxt Thm.theoremK)
   249   |> map (Element.pretty_statement ctxt Thm.theoremK)
   250   |> Pretty.chunks2 |> Pretty.string_of;
   250   |> Pretty.chunks2;
   251 
   251 
   252 fun string_of_thms ctxt args =
   252 fun pretty_thms ctxt args =
   253   Pretty.string_of (Proof_Context.pretty_fact ctxt ("", Attrib.eval_thms ctxt args));
   253   Proof_Context.pretty_fact ctxt ("", Attrib.eval_thms ctxt args);
   254 
   254 
   255 fun string_of_prfs full state arg =
   255 fun pretty_prfs full state arg =
   256   Pretty.string_of
   256   (case arg of
   257     (case arg of
   257     NONE =>
   258       NONE =>
   258       let
   259         let
   259         val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state);
   260           val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state);
   260         val thy = Proof_Context.theory_of ctxt;
   261           val thy = Proof_Context.theory_of ctxt;
   261         val prf = Thm.proof_of thm;
   262           val prf = Thm.proof_of thm;
   262         val prop = Thm.full_prop_of thm;
   263           val prop = Thm.full_prop_of thm;
   263         val prf' = Proofterm.rewrite_proof_notypes ([], []) prf;
   264           val prf' = Proofterm.rewrite_proof_notypes ([], []) prf;
   264       in
   265         in
   265         Proof_Syntax.pretty_proof ctxt
   266           Proof_Syntax.pretty_proof ctxt
   266           (if full then Proofterm.reconstruct_proof thy prop prf' else prf')
   267             (if full then Proofterm.reconstruct_proof thy prop prf' else prf')
   267       end
   268         end
   268   | SOME srcs =>
   269     | SOME srcs =>
   269       let
   270         let
   270         val ctxt = Toplevel.context_of state;
   271           val ctxt = Toplevel.context_of state;
   271         val pretty_proof = Proof_Syntax.pretty_standard_proof_of ctxt full;
   272           val pretty_proof = Proof_Syntax.pretty_standard_proof_of ctxt full;
   272       in Pretty.chunks (map pretty_proof (Attrib.eval_thms ctxt srcs)) end);
   273         in Pretty.chunks (map pretty_proof (Attrib.eval_thms ctxt srcs)) end);
   273 
   274 
   274 fun pretty_prop ctxt s =
   275 fun string_of_prop ctxt s =
       
   276   let
   275   let
   277     val prop = Syntax.read_prop ctxt s;
   276     val prop = Syntax.read_prop ctxt s;
   278     val ctxt' = Proof_Context.augment prop ctxt;
   277     val ctxt' = Proof_Context.augment prop ctxt;
   279   in Pretty.string_of (Pretty.quote (Syntax.pretty_term ctxt' prop)) end;
   278   in Pretty.quote (Syntax.pretty_term ctxt' prop) end;
   280 
   279 
   281 fun string_of_term ctxt s =
   280 fun pretty_term ctxt s =
   282   let
   281   let
   283     val t = Syntax.read_term ctxt s;
   282     val t = Syntax.read_term ctxt s;
   284     val T = Term.type_of t;
   283     val T = Term.type_of t;
   285     val ctxt' = Proof_Context.augment t ctxt;
   284     val ctxt' = Proof_Context.augment t ctxt;
   286   in
   285   in
   287     Pretty.string_of
   286     Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk,
   288       (Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk,
   287       Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)]
   289         Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)])
       
   290   end;
   288   end;
   291 
   289 
   292 fun string_of_type ctxt (s, NONE) =
   290 fun pretty_type ctxt (s, NONE) =
   293       let val T = Syntax.read_typ ctxt s
   291       let val T = Syntax.read_typ ctxt s
   294       in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt T)) end
   292       in Pretty.quote (Syntax.pretty_typ ctxt T) end
   295   | string_of_type ctxt (s1, SOME s2) =
   293   | pretty_type ctxt (s1, SOME s2) =
   296       let
   294       let
   297         val ctxt' = Config.put show_sorts true ctxt;
   295         val ctxt' = Config.put show_sorts true ctxt;
   298         val raw_T = Syntax.parse_typ ctxt' s1;
   296         val raw_T = Syntax.parse_typ ctxt' s1;
   299         val S = Syntax.read_sort ctxt' s2;
   297         val S = Syntax.read_sort ctxt' s2;
   300         val T =
   298         val T =
   301           Syntax.check_term ctxt'
   299           Syntax.check_term ctxt'
   302             (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S)))
   300             (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S)))
   303           |> Logic.dest_type;
   301           |> Logic.dest_type;
   304       in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt' T)) end;
   302       in Pretty.quote (Syntax.pretty_typ ctxt' T) end;
   305 
   303 
   306 fun print_item string_of (modes, arg) = Toplevel.keep (fn state =>
   304 fun print_item pretty (modes, arg) = Toplevel.keep (fn state =>
   307   Print_Mode.with_modes modes (fn () => writeln (string_of state arg)) ());
   305   Print_Mode.with_modes modes (fn () => Pretty.writeln (pretty state arg)) ());
   308 
   306 
   309 in
   307 in
   310 
   308 
   311 val print_stmts = print_item (string_of_stmts o Toplevel.context_of);
   309 val print_stmts = print_item (pretty_stmts o Toplevel.context_of);
   312 val print_thms = print_item (string_of_thms o Toplevel.context_of);
   310 val print_thms = print_item (pretty_thms o Toplevel.context_of);
   313 val print_prfs = print_item o string_of_prfs;
   311 val print_prfs = print_item o pretty_prfs;
   314 val print_prop = print_item (string_of_prop o Toplevel.context_of);
   312 val print_prop = print_item (pretty_prop o Toplevel.context_of);
   315 val print_term = print_item (string_of_term o Toplevel.context_of);
   313 val print_term = print_item (pretty_term o Toplevel.context_of);
   316 val print_type = print_item (string_of_type o Toplevel.context_of);
   314 val print_type = print_item (pretty_type o Toplevel.context_of);
   317 
   315 
   318 end;
   316 end;
   319 
   317 
   320 end;
   318 end;