src/Pure/sign.ML
changeset 39133 70d3915c92f0
parent 37145 01aa36932739
child 39137 ccb53edd59f0
equal deleted inserted replaced
39132:ba17ca3acdd3 39133:70d3915c92f0
    65   val certify_typ_mode: Type.mode -> theory -> typ -> typ
    65   val certify_typ_mode: Type.mode -> theory -> typ -> typ
    66   val certify': bool -> Pretty.pp -> bool -> Consts.T -> theory -> term -> term * typ * int
    66   val certify': bool -> Pretty.pp -> bool -> Consts.T -> theory -> term -> term * typ * int
    67   val certify_term: theory -> term -> term * typ * int
    67   val certify_term: theory -> term -> term * typ * int
    68   val cert_term: theory -> term -> term
    68   val cert_term: theory -> term -> term
    69   val cert_prop: theory -> term -> term
    69   val cert_prop: theory -> term -> term
    70   val no_frees: Pretty.pp -> term -> term
    70   val no_frees: Proof.context -> term -> term
    71   val no_vars: Pretty.pp -> term -> term
    71   val no_vars: Proof.context -> term -> term
    72   val add_types: (binding * int * mixfix) list -> theory -> theory
    72   val add_types: (binding * int * mixfix) list -> theory -> theory
    73   val add_nonterminals: binding list -> theory -> theory
    73   val add_nonterminals: binding list -> theory -> theory
    74   val add_type_abbrev: binding * string list * typ -> theory -> theory
    74   val add_type_abbrev: binding * string list * typ -> theory -> theory
    75   val add_syntax: (string * string * mixfix) list -> theory -> theory
    75   val add_syntax: (string * string * mixfix) list -> theory -> theory
    76   val add_syntax_i: (string * typ * mixfix) list -> theory -> theory
    76   val add_syntax_i: (string * typ * mixfix) list -> theory -> theory
   318 end;
   318 end;
   319 
   319 
   320 
   320 
   321 (* specifications *)
   321 (* specifications *)
   322 
   322 
   323 fun no_variables kind add addT mk mkT pp tm =
   323 fun no_variables kind add addT mk mkT ctxt tm =
   324   (case (add tm [], addT tm []) of
   324   (case (add tm [], addT tm []) of
   325     ([], []) => tm
   325     ([], []) => tm
   326   | (frees, tfrees) => error (Pretty.string_of (Pretty.block
   326   | (frees, tfrees) => error (Pretty.string_of (Pretty.block
   327       (Pretty.str ("Illegal " ^ kind ^ " variable(s) in term:") :: Pretty.brk 1 ::
   327       (Pretty.str ("Illegal " ^ kind ^ " variable(s) in term:") :: Pretty.brk 1 ::
   328        Pretty.commas (map (Pretty.term pp o mk) frees @ map (Pretty.typ pp o mkT) tfrees)))));
   328        Pretty.commas
       
   329         (map (Syntax.pretty_term ctxt o mk) frees @ map (Syntax.pretty_typ ctxt o mkT) tfrees)))));
   329 
   330 
   330 val no_frees = no_variables "free" Term.add_frees Term.add_tfrees Free TFree;
   331 val no_frees = no_variables "free" Term.add_frees Term.add_tfrees Free TFree;
   331 val no_vars = no_variables "schematic" Term.add_vars Term.add_tvars Var TVar;
   332 val no_vars = no_variables "schematic" Term.add_vars Term.add_tvars Var TVar;
   332 
   333 
   333 
   334 
   432 
   433 
   433 (* abbreviations *)
   434 (* abbreviations *)
   434 
   435 
   435 fun add_abbrev mode (b, raw_t) thy =
   436 fun add_abbrev mode (b, raw_t) thy =
   436   let
   437   let
   437     val pp = Syntax.pp_global thy;
   438     val ctxt = Syntax.init_pretty_global thy;
   438     val prep_tm = no_frees pp o Term.no_dummy_patterns o cert_term_abbrev thy;
   439     val prep_tm = no_frees ctxt o Term.no_dummy_patterns o cert_term_abbrev thy;
   439     val t = (prep_tm raw_t handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg)
   440     val t = (prep_tm raw_t handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg)
   440       handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.str_of b));
   441       handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.str_of b));
   441     val (res, consts') = consts_of thy
   442     val (res, consts') = consts_of thy
   442       |> Consts.abbreviate pp (tsig_of thy) (naming_of thy) mode (b, t);
   443       |> Consts.abbreviate (Syntax.pp ctxt) (tsig_of thy) (naming_of thy) mode (b, t);
   443   in (res, thy |> map_consts (K consts')) end;
   444   in (res, thy |> map_consts (K consts')) end;
   444 
   445 
   445 fun revert_abbrev mode c = map_consts (Consts.revert_abbrev mode c);
   446 fun revert_abbrev mode c = map_consts (Consts.revert_abbrev mode c);
   446 
   447 
   447 
   448