src/Pure/Isar/proof_context.ML
author wenzelm
Sat, 21 Apr 2007 11:07:42 +0200
changeset 22763 5c1752279f25
parent 22728 ecbbdf50df2f
child 22773 9bb135fa5206
permissions -rw-r--r--
added decode_term (belongs to Syntax module);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/proof_context.ML
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     4
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
     5
The key concept of Isar proof contexts: elevates primitive local
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
     6
reasoning Gamma |- phi to a structured concept, with generic context
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
     7
elements.  See also structure Variable and Assumption.
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     8
*)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
     9
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
    10
signature PROOF_CONTEXT =
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
    11
sig
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    12
  val theory_of: Proof.context -> theory
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    13
  val init: theory -> Proof.context
21667
ce813b82c88b add_notation: permissive about undeclared consts;
wenzelm
parents: 21648
diff changeset
    14
  val is_stmt: Proof.context -> bool
ce813b82c88b add_notation: permissive about undeclared consts;
wenzelm
parents: 21648
diff changeset
    15
  val set_stmt: bool -> Proof.context -> Proof.context
ce813b82c88b add_notation: permissive about undeclared consts;
wenzelm
parents: 21648
diff changeset
    16
  val restore_stmt: Proof.context -> Proof.context -> Proof.context
ce813b82c88b add_notation: permissive about undeclared consts;
wenzelm
parents: 21648
diff changeset
    17
  val naming_of: Proof.context -> NameSpace.naming
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    18
  val full_name: Proof.context -> bstring -> string
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    19
  val consts_of: Proof.context -> Consts.T
21183
a76f457b6d86 added const_syntax_name;
wenzelm
parents: 20784
diff changeset
    20
  val const_syntax_name: Proof.context -> string -> string
20784
eece9aaaf352 Syntax.mode;
wenzelm
parents: 20631
diff changeset
    21
  val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    22
  val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
22358
4d8a9e3a29f8 added theorems_of;
wenzelm
parents: 22352
diff changeset
    23
  val theorems_of: Proof.context -> thm list NameSpace.table
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    24
  val fact_index_of: Proof.context -> FactIndex.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    25
  val transfer: theory -> Proof.context -> Proof.context
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
    26
  val theory: (theory -> theory) -> Proof.context -> Proof.context
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
    27
  val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    28
  val pretty_term: Proof.context -> term -> Pretty.T
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
    29
  val pretty_term_abbrev: Proof.context -> term -> Pretty.T
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    30
  val pretty_typ: Proof.context -> typ -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    31
  val pretty_sort: Proof.context -> sort -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    32
  val pretty_classrel: Proof.context -> class list -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    33
  val pretty_arity: Proof.context -> arity -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    34
  val pp: Proof.context -> Pretty.pp
20253
636ac45d100f rename legacy_pretty_thm to pretty_thm_legacy;
wenzelm
parents: 20244
diff changeset
    35
  val pretty_thm_legacy: bool -> thm -> Pretty.T
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    36
  val pretty_thm: Proof.context -> thm -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    37
  val pretty_thms: Proof.context -> thm list -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    38
  val pretty_fact: Proof.context -> string * thm list -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    39
  val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    40
  val pretty_proof_of: Proof.context -> bool -> thm -> Pretty.T
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    41
  val string_of_typ: Proof.context -> typ -> string
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    42
  val string_of_term: Proof.context -> term -> string
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    43
  val string_of_thm: Proof.context -> thm -> string
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    44
  val read_typ: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    45
  val read_typ_syntax: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    46
  val read_typ_abbrev: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    47
  val cert_typ: Proof.context -> typ -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    48
  val cert_typ_syntax: Proof.context -> typ -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    49
  val cert_typ_abbrev: Proof.context -> typ -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    50
  val get_skolem: Proof.context -> string -> string
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    51
  val revert_skolem: Proof.context -> string -> string
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    52
  val revert_skolems: Proof.context -> term -> term
22763
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
    53
  val decode_term: Proof.context -> term -> term
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    54
  val read_termTs: Proof.context -> (string -> bool) -> (indexname -> typ option)
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
    55
    -> (indexname -> sort option) -> string list -> (string * typ) list
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
    56
    -> term list * (indexname * typ) list
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    57
  val read_termTs_schematic: Proof.context -> (string -> bool) -> (indexname -> typ option)
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
    58
    -> (indexname -> sort option) -> string list -> (string * typ) list
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
    59
    -> term list * (indexname * typ) list
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
    60
  val read_prop_legacy: Proof.context -> string -> term
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    61
  val read_term: Proof.context -> string -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    62
  val read_prop: Proof.context -> string -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    63
  val read_prop_schematic: Proof.context -> string -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    64
  val read_term_pats: typ -> Proof.context -> string list -> term list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    65
  val read_prop_pats: Proof.context -> string list -> term list
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
    66
  val read_term_abbrev: Proof.context -> string -> term
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    67
  val cert_term: Proof.context -> term -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    68
  val cert_prop: Proof.context -> term -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    69
  val cert_term_pats: typ -> Proof.context -> term list -> term list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    70
  val cert_prop_pats: Proof.context -> term list -> term list
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
    71
  val infer_types: Proof.context -> term list -> term list
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
    72
  val infer_types_pats: Proof.context -> term list -> term list
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    73
  val infer_type: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    74
  val inferred_param: string -> Proof.context -> (string * typ) * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    75
  val inferred_fixes: Proof.context -> (string * typ) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    76
  val read_tyname: Proof.context -> string -> typ
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    77
  val read_const: Proof.context -> string -> term
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    78
  val goal_export: Proof.context -> Proof.context -> thm list -> thm list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    79
  val export: Proof.context -> Proof.context -> thm list -> thm list
21531
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
    80
  val export_morphism: Proof.context -> Proof.context -> morphism
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    81
  val add_binds: (indexname * string option) list -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    82
  val add_binds_i: (indexname * term option) list -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    83
  val auto_bind_goal: term list -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    84
  val auto_bind_facts: term list -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    85
  val match_bind: bool -> (string list * string) list -> Proof.context -> term list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    86
  val match_bind_i: bool -> (term list * term) list -> Proof.context -> term list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    87
  val read_propp: Proof.context * (string * string list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    88
    -> Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    89
  val cert_propp: Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    90
    -> Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    91
  val read_propp_schematic: Proof.context * (string * string list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    92
    -> Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    93
  val cert_propp_schematic: Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    94
    -> Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    95
  val bind_propp: Proof.context * (string * string list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    96
    -> Proof.context * (term list list * (Proof.context -> Proof.context))
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    97
  val bind_propp_i: Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    98
    -> Proof.context * (term list list * (Proof.context -> Proof.context))
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
    99
  val bind_propp_schematic: Proof.context * (string * string list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   100
    -> Proof.context * (term list list * (Proof.context -> Proof.context))
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   101
  val bind_propp_schematic_i: Proof.context * (term * term list) list list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   102
    -> Proof.context * (term list list * (Proof.context -> Proof.context))
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   103
  val fact_tac: thm list -> int -> tactic
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   104
  val some_fact_tac: Proof.context -> int -> tactic
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   105
  val get_thm: Proof.context -> thmref -> thm
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   106
  val get_thm_closure: Proof.context -> thmref -> thm
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   107
  val get_thms: Proof.context -> thmref -> thm list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   108
  val get_thms_closure: Proof.context -> thmref -> thm list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   109
  val valid_thms: Proof.context -> string * thm list -> bool
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   110
  val lthms_containing: Proof.context -> FactIndex.spec -> (string * thm list) list
22352
f15118a79c0e add_path for naming in proof contexts
haftmann
parents: 21824
diff changeset
   111
  val add_path: string -> Proof.context -> Proof.context
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   112
  val no_base_names: Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   113
  val qualified_names: Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   114
  val sticky_prefix: string -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   115
  val restore_naming: Proof.context -> Proof.context -> Proof.context
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   116
  val reset_naming: Proof.context -> Proof.context
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   117
  val hide_thms: bool -> string list -> Proof.context -> Proof.context
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   118
  val put_thms: string * thm list option -> Proof.context -> Proof.context
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   119
  val note_thmss: string ->
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   120
    ((bstring * attribute list) * (thmref * attribute list) list) list ->
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   121
      Proof.context -> (bstring * thm list) list * Proof.context
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   122
  val note_thmss_i: string ->
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   123
    ((bstring * attribute list) * (thm list * attribute list) list) list ->
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   124
      Proof.context -> (bstring * thm list) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   125
  val read_vars: (string * string option * mixfix) list -> Proof.context ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   126
    (string * typ option * mixfix) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   127
  val cert_vars: (string * typ option * mixfix) list -> Proof.context ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   128
    (string * typ option * mixfix) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   129
  val read_vars_legacy: (string * string option * mixfix) list -> Proof.context ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   130
    (string * typ option * mixfix) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   131
  val cert_vars_legacy: (string * typ option * mixfix) list -> Proof.context ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   132
    (string * typ option * mixfix) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   133
  val add_fixes: (string * string option * mixfix) list ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   134
    Proof.context -> string list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   135
  val add_fixes_i: (string * typ option * mixfix) list ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   136
    Proof.context -> string list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   137
  val add_fixes_legacy: (string * typ option * mixfix) list ->
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   138
    Proof.context -> string list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   139
  val auto_fixes: Proof.context * (term list list * 'a) -> Proof.context * (term list list * 'a)
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   140
  val bind_fixes: string list -> Proof.context -> (term -> term) * Proof.context
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   141
  val add_assms: Assumption.export ->
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   142
    ((string * attribute list) * (string * string list) list) list ->
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   143
    Proof.context -> (bstring * thm list) list * Proof.context
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   144
  val add_assms_i: Assumption.export ->
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   145
    ((string * attribute list) * (term * term list) list) list ->
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   146
    Proof.context -> (bstring * thm list) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   147
  val add_cases: bool -> (string * RuleCases.T option) list -> Proof.context -> Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   148
  val apply_case: RuleCases.T -> Proof.context -> (string * term list) list * Proof.context
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   149
  val get_case: Proof.context -> string -> string option list -> RuleCases.T
21208
62ccdaf9369a replaced const_syntax by notation, which operates on terms;
wenzelm
parents: 21183
diff changeset
   150
  val add_notation: Syntax.mode -> (term * mixfix) list ->
62ccdaf9369a replaced const_syntax by notation, which operates on terms;
wenzelm
parents: 21183
diff changeset
   151
    Proof.context -> Proof.context
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   152
  val target_notation: Syntax.mode -> (term * mixfix) list -> morphism ->
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   153
    Context.generic -> Context.generic
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   154
  val set_expand_abbrevs: bool -> Proof.context -> Proof.context
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   155
  val add_abbrev: string -> bstring * term -> Proof.context -> (term * term) * Proof.context
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   156
  val target_abbrev: Syntax.mode -> (string * mixfix) * term -> morphism ->
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   157
    Context.generic -> Context.generic
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   158
  val local_abbrev: string * term -> Proof.context -> (term * term) * Proof.context
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
   159
  val verbose: bool ref
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
   160
  val setmp_verbose: ('a -> 'b) -> 'a -> 'b
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   161
  val print_syntax: Proof.context -> unit
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   162
  val print_abbrevs: Proof.context -> unit
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   163
  val print_binds: Proof.context -> unit
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   164
  val print_lthms: Proof.context -> unit
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   165
  val print_cases: Proof.context -> unit
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   166
  val debug: bool ref
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
   167
  val prems_limit: int ref
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   168
  val pretty_ctxt: Proof.context -> Pretty.T list
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   169
  val pretty_context: Proof.context -> Pretty.T list
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   170
end;
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   171
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   172
structure ProofContext: PROOF_CONTEXT =
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   173
struct
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   174
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   175
val theory_of = Context.theory_of_proof;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   176
val tsig_of = Sign.tsig_of o theory_of;
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   177
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   178
val init = Context.init_proof;
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
   179
7270
2b64729d9acb warn_vars;
wenzelm
parents: 7200
diff changeset
   180
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   181
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   182
(** Isar proof context information **)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   183
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   184
datatype ctxt =
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   185
  Ctxt of
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   186
   {is_stmt: bool,                                (*inner statement mode*)
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   187
    naming: NameSpace.naming,                     (*local naming conventions*)
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   188
    syntax: LocalSyntax.T,                        (*local syntax*)
19033
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   189
    consts: Consts.T * Consts.T,                  (*global/local consts*)
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   190
    thms: thm list NameSpace.table * FactIndex.T, (*local thms*)
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   191
    cases: (string * (RuleCases.T * bool)) list}; (*local contexts*)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   192
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   193
fun make_ctxt (is_stmt, naming, syntax, consts, thms, cases) =
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   194
  Ctxt {is_stmt = is_stmt, naming = naming, syntax = syntax,
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   195
    consts = consts, thms = thms, cases = cases};
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   196
19079
9a7678a0736d added put_thms_internal: local_naming, no fact index;
wenzelm
parents: 19062
diff changeset
   197
val local_naming = NameSpace.default_naming |> NameSpace.add_path "local";
9a7678a0736d added put_thms_internal: local_naming, no fact index;
wenzelm
parents: 19062
diff changeset
   198
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   199
structure ContextData = ProofDataFun
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   200
(
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   201
  val name = "Isar/context";
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   202
  type T = ctxt;
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   203
  fun init thy =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   204
    make_ctxt (false, local_naming, LocalSyntax.init thy,
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   205
      (Sign.consts_of thy, Sign.consts_of thy),
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   206
      (NameSpace.empty_table, FactIndex.empty), []);
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   207
  fun print _ _ = ();
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   208
);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   209
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18699
diff changeset
   210
val _ = Context.add_setup ContextData.init;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   211
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   212
fun rep_context ctxt = ContextData.get ctxt |> (fn Ctxt args => args);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   213
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   214
fun map_context f =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   215
  ContextData.map (fn Ctxt {is_stmt, naming, syntax, consts, thms, cases} =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   216
    make_ctxt (f (is_stmt, naming, syntax, consts, thms, cases)));
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   217
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   218
fun set_stmt b =
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   219
  map_context (fn (_, naming, syntax, consts, thms, cases) =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   220
    (b, naming, syntax, consts, thms, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   221
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   222
fun map_naming f =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   223
  map_context (fn (is_stmt, naming, syntax, consts, thms, cases) =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   224
    (is_stmt, f naming, syntax, consts, thms, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   225
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   226
fun map_syntax f =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   227
  map_context (fn (is_stmt, naming, syntax, consts, thms, cases) =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   228
    (is_stmt, naming, f syntax, consts, thms, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   229
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   230
fun map_consts f =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   231
  map_context (fn (is_stmt, naming, syntax, consts, thms, cases) =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   232
    (is_stmt, naming, syntax, f consts, thms, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   233
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   234
fun map_thms f =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   235
  map_context (fn (is_stmt, naming, syntax, consts, thms, cases) =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   236
    (is_stmt, naming, syntax, consts, f thms, cases));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   237
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   238
fun map_cases f =
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   239
  map_context (fn (is_stmt, naming, syntax, consts, thms, cases) =>
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   240
    (is_stmt, naming, syntax, consts, thms, f cases));
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   241
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   242
val is_stmt = #is_stmt o rep_context;
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   243
fun restore_stmt ctxt = set_stmt (is_stmt ctxt);
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   244
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   245
val naming_of = #naming o rep_context;
19387
6af442fa80c3 added full_name;
wenzelm
parents: 19371
diff changeset
   246
val full_name = NameSpace.full o naming_of;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   247
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   248
val syntax_of = #syntax o rep_context;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   249
val syn_of = LocalSyntax.syn_of o syntax_of;
19543
e1b81ecd4580 added set_syntax_mode, restore_syntax_mode;
wenzelm
parents: 19482
diff changeset
   250
val set_syntax_mode = map_syntax o LocalSyntax.set_mode;
e1b81ecd4580 added set_syntax_mode, restore_syntax_mode;
wenzelm
parents: 19482
diff changeset
   251
val restore_syntax_mode = map_syntax o LocalSyntax.restore_mode o syntax_of;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   252
19033
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   253
val consts_of = #2 o #consts o rep_context;
21183
a76f457b6d86 added const_syntax_name;
wenzelm
parents: 20784
diff changeset
   254
val const_syntax_name = Consts.syntax_name o consts_of;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   255
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   256
val thms_of = #thms o rep_context;
22358
4d8a9e3a29f8 added theorems_of;
wenzelm
parents: 22352
diff changeset
   257
val theorems_of = #1 o thms_of;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   258
val fact_index_of = #2 o thms_of;
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   259
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   260
val cases_of = #cases o rep_context;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   261
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   262
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   263
(* theory transfer *)
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
   264
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   265
fun transfer_syntax thy =
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   266
  map_syntax (LocalSyntax.rebuild thy) #>
19033
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   267
  map_consts (fn consts as (global_consts, local_consts) =>
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   268
    let val thy_consts = Sign.consts_of thy in
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   269
      if Consts.eq_consts (thy_consts, global_consts) then consts
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   270
      else (thy_consts, Consts.merge (thy_consts, local_consts))
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
   271
    end);
17072
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   272
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   273
fun transfer thy = Context.transfer_proof thy #> transfer_syntax thy;
17072
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   274
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   275
fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   276
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   277
fun theory_result f ctxt =
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   278
  let val (res, thy') = f (theory_of ctxt)
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
   279
  in (res, ctxt |> transfer thy') end;
19019
412009243085 added map_theory;
wenzelm
parents: 19001
diff changeset
   280
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
   281
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
   282
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   283
(** pretty printing **)
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   284
19310
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   285
local
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   286
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   287
fun rewrite_term warn thy rews t =
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   288
  if can Term.type_of t then Pattern.rewrite_term thy rews [] t
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   289
  else (if warn then warning "Printing malformed term -- cannot expand abbreviations" else (); t);
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   290
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   291
fun pretty_term' abbrevs ctxt t =
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   292
  let
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   293
    val thy = theory_of ctxt;
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   294
    val syntax = syntax_of ctxt;
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   295
    val consts = consts_of ctxt;
22587
5454b06320fb renamed Output.has_mode to print_mode_active;
wenzelm
parents: 22358
diff changeset
   296
    val do_abbrevs = abbrevs andalso not (print_mode_active "no_abbrevs");
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   297
    val t' = t
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   298
      |> do_abbrevs ? rewrite_term true thy (Consts.abbrevs_of consts (! print_mode @ [""]))
21824
153fad1e7318 target_abbrev: internal mode for abbrevs;
wenzelm
parents: 21814
diff changeset
   299
      |> do_abbrevs ? rewrite_term false thy (Consts.abbrevs_of consts [Syntax.internalM])
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   300
      |> Sign.extern_term (Consts.extern_early consts) thy
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   301
      |> LocalSyntax.extern_term syntax;
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21752
diff changeset
   302
  in Sign.pretty_term' ctxt (LocalSyntax.syn_of syntax) (Consts.extern consts) t' end;
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   303
19310
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   304
in
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   305
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   306
val pretty_term = pretty_term' true;
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   307
val pretty_term_abbrev = pretty_term' false;
19310
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   308
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   309
end;
9b2080d9ed28 subtract (op =);
wenzelm
parents: 19274
diff changeset
   310
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16348
diff changeset
   311
fun pretty_typ ctxt T = Sign.pretty_typ (theory_of ctxt) T;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16348
diff changeset
   312
fun pretty_sort ctxt S = Sign.pretty_sort (theory_of ctxt) S;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16348
diff changeset
   313
fun pretty_classrel ctxt cs = Sign.pretty_classrel (theory_of ctxt) cs;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16348
diff changeset
   314
fun pretty_arity ctxt ar = Sign.pretty_arity (theory_of ctxt) ar;
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   315
14974
b1ecb7859c99 avoid premature evaluation of syn_of (wastes time in conjunction with pp);
wenzelm
parents: 14901
diff changeset
   316
fun pp ctxt = Pretty.pp (pretty_term ctxt, pretty_typ ctxt, pretty_sort ctxt,
b1ecb7859c99 avoid premature evaluation of syn_of (wastes time in conjunction with pp);
wenzelm
parents: 14901
diff changeset
   317
  pretty_classrel ctxt, pretty_arity ctxt);
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   318
20253
636ac45d100f rename legacy_pretty_thm to pretty_thm_legacy;
wenzelm
parents: 20244
diff changeset
   319
fun pretty_thm_legacy quote th =
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   320
  let
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   321
    val thy = Thm.theory_of_thm th;
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   322
    val pp =
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   323
      if Theory.eq_thy (thy, ProtoPure.thy) then Sign.pp thy
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   324
      else pp (init thy);
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   325
  in Display.pretty_thm_aux pp quote false [] th end;
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
   326
17451
cfa8b1ebfc9a added auto_fix (from proof.ML);
wenzelm
parents: 17412
diff changeset
   327
fun pretty_thm ctxt th =
20575
6b1c69265331 pretty_thm: graceful treatment of ProtoPure.thy;
wenzelm
parents: 20367
diff changeset
   328
  let
6b1c69265331 pretty_thm: graceful treatment of ProtoPure.thy;
wenzelm
parents: 20367
diff changeset
   329
    val thy = theory_of ctxt;
6b1c69265331 pretty_thm: graceful treatment of ProtoPure.thy;
wenzelm
parents: 20367
diff changeset
   330
    val (pp, asms) =
6b1c69265331 pretty_thm: graceful treatment of ProtoPure.thy;
wenzelm
parents: 20367
diff changeset
   331
      if Theory.eq_thy (thy, ProtoPure.thy) then (Sign.pp thy, [])
21583
6fb553dc039b removed export_standard_morphism;
wenzelm
parents: 21569
diff changeset
   332
      else (pp ctxt, map Thm.term_of (Assumption.assms_of ctxt));
20575
6b1c69265331 pretty_thm: graceful treatment of ProtoPure.thy;
wenzelm
parents: 20367
diff changeset
   333
  in Display.pretty_thm_aux pp false true asms th end;
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   334
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   335
fun pretty_thms ctxt [th] = pretty_thm ctxt th
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   336
  | pretty_thms ctxt ths = Pretty.blk (0, Pretty.fbreaks (map (pretty_thm ctxt) ths));
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   337
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   338
fun pretty_fact ctxt ("", ths) = pretty_thms ctxt ths
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   339
  | pretty_fact ctxt (a, [th]) =
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   340
      Pretty.block [Pretty.str (a ^ ":"), Pretty.brk 1, pretty_thm ctxt th]
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   341
  | pretty_fact ctxt (a, ths) =
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   342
      Pretty.block (Pretty.fbreaks (Pretty.str (a ^ ":") :: map (pretty_thm ctxt) ths));
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   343
17072
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   344
fun pretty_proof ctxt prf =
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   345
  pretty_term_abbrev (ctxt |> transfer_syntax (ProofSyntax.proof_syntax prf (theory_of ctxt)))
17072
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   346
    (ProofSyntax.term_of_proof prf);
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   347
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   348
fun pretty_proof_of ctxt full th =
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   349
  pretty_proof ctxt (ProofSyntax.proof_of full th);
501c28052aea added transfer;
wenzelm
parents: 16992
diff changeset
   350
17860
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   351
val string_of_typ = Pretty.string_of oo pretty_typ;
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   352
val string_of_term = Pretty.string_of oo pretty_term;
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   353
val string_of_thm = Pretty.string_of oo pretty_thm;
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   354
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   355
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   356
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   357
(** prepare types **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   358
9504
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   359
local
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   360
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   361
fun read_typ_aux read ctxt s =
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21752
diff changeset
   362
  read (syn_of ctxt) ctxt (Variable.def_sort ctxt) s;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   363
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   364
fun cert_typ_aux cert ctxt raw_T =
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   365
  cert (theory_of ctxt) raw_T handle TYPE (msg, _, _) => error msg;
9504
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   366
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   367
in
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   368
16348
7504fe04170f renamed extern to extern_thm;
wenzelm
parents: 16147
diff changeset
   369
val read_typ        = read_typ_aux Sign.read_typ';
7504fe04170f renamed extern to extern_thm;
wenzelm
parents: 16147
diff changeset
   370
val read_typ_syntax = read_typ_aux Sign.read_typ_syntax';
7504fe04170f renamed extern to extern_thm;
wenzelm
parents: 16147
diff changeset
   371
val read_typ_abbrev = read_typ_aux Sign.read_typ_abbrev';
7504fe04170f renamed extern to extern_thm;
wenzelm
parents: 16147
diff changeset
   372
val cert_typ        = cert_typ_aux Sign.certify_typ;
7504fe04170f renamed extern to extern_thm;
wenzelm
parents: 16147
diff changeset
   373
val cert_typ_syntax = cert_typ_aux Sign.certify_typ_syntax;
7504fe04170f renamed extern to extern_thm;
wenzelm
parents: 16147
diff changeset
   374
val cert_typ_abbrev = cert_typ_aux Sign.certify_typ_abbrev;
9504
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   375
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   376
end;
8168600e88a5 typ_no_norm;
wenzelm
parents: 9470
diff changeset
   377
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   378
7679
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   379
(* internalize Skolem constants *)
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   380
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   381
val lookup_skolem = AList.lookup (op =) o Variable.fixes_of;
18187
ec44df8ffd21 added revert_skolem, mk_def, add_def;
wenzelm
parents: 18152
diff changeset
   382
fun get_skolem ctxt x = the_default x (lookup_skolem ctxt x);
7679
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   383
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   384
fun no_skolem internal x =
20086
94ca946fb689 adapted to more efficient Name/Variable implementation;
wenzelm
parents: 20049
diff changeset
   385
  if can Name.dest_skolem x then
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   386
    error ("Illegal reference to internal Skolem constant: " ^ quote x)
20086
94ca946fb689 adapted to more efficient Name/Variable implementation;
wenzelm
parents: 20049
diff changeset
   387
  else if not internal andalso can Name.dest_internal x then
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   388
    error ("Illegal reference to internal variable: " ^ quote x)
7679
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   389
  else x;
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   390
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   391
20244
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   392
(* revert Skolem constants -- approximation only! *)
18255
5ef79b75a002 revert_skolem: fall back on Syntax.deskolem;
wenzelm
parents: 18211
diff changeset
   393
20244
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   394
fun revert_skolem ctxt =
9133
bc3742f62b80 added extern_skolem;
wenzelm
parents: 9007
diff changeset
   395
  let
20244
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   396
    val rev_fixes = map Library.swap (Variable.fixes_of ctxt);
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   397
    val revert = AList.lookup (op =) rev_fixes;
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   398
  in
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   399
    fn x =>
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   400
      (case revert x of
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   401
        SOME x' => x'
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   402
      | NONE => perhaps (try Name.dest_skolem) x)
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   403
   end;
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   404
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   405
fun revert_skolems ctxt =
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   406
  let
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   407
    val revert = revert_skolem ctxt;
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   408
    fun reverts (Free (x, T)) = Free (revert x, T)
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   409
      | reverts (t $ u) = reverts t $ reverts u
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   410
      | reverts (Abs (x, T, t)) = Abs (x, T, reverts t)
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   411
      | reverts a = a;
89a407400874 replaced extern_skolem by slightly more simplistic revert_skolems;
wenzelm
parents: 20234
diff changeset
   412
  in reverts end
9133
bc3742f62b80 added extern_skolem;
wenzelm
parents: 9007
diff changeset
   413
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   414
18187
ec44df8ffd21 added revert_skolem, mk_def, add_def;
wenzelm
parents: 18152
diff changeset
   415
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   416
(** prepare terms and propositions **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   417
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   418
(*
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   419
  (1) read / certify wrt. theory of context
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   420
  (2) intern Skolem constants
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   421
  (3) expand term bindings
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   422
*)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   423
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   424
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   425
(* read wrt. theory *)     (*exception ERROR*)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   426
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   427
fun read_def_termTs freeze pp syn ctxt (types, sorts, used) fixed sTs =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   428
  Sign.read_def_terms' pp (Sign.is_logtype (theory_of ctxt)) syn (consts_of ctxt) fixed
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21752
diff changeset
   429
    ctxt (types, sorts) used freeze sTs;
5874
a58d4528121d renamed init_context to init;
wenzelm
parents: 5819
diff changeset
   430
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   431
fun read_def_termT freeze pp syn ctxt defaults fixed sT =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   432
  apfst hd (read_def_termTs freeze pp syn ctxt defaults fixed [sT]);
14828
15d12761ba54 improved output; refer to Pretty.pp;
wenzelm
parents: 14780
diff changeset
   433
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   434
fun read_term_thy freeze pp syn thy defaults fixed s =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   435
  #1 (read_def_termT freeze pp syn thy defaults fixed (s, dummyT));
5874
a58d4528121d renamed init_context to init;
wenzelm
parents: 5819
diff changeset
   436
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   437
fun read_prop_thy freeze pp syn thy defaults fixed s =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   438
  #1 (read_def_termT freeze pp syn thy defaults fixed (s, propT));
12072
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
   439
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   440
fun read_terms_thy freeze pp syn thy defaults fixed =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   441
  #1 o read_def_termTs freeze pp syn thy defaults fixed o map (rpair dummyT);
12072
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
   442
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   443
fun read_props_thy freeze pp syn thy defaults fixed =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   444
  #1 o read_def_termTs freeze pp syn thy defaults fixed o map (rpair propT);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   445
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   446
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   447
(* local abbreviations *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   448
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   449
fun certify_consts ctxt =
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   450
  Consts.certify (pp ctxt) (tsig_of ctxt) (consts_of ctxt);
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   451
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   452
fun reject_schematic (Var (xi, _)) =
22678
23963361278c Term.string_of_vname;
wenzelm
parents: 22670
diff changeset
   453
      error ("Unbound schematic variable: " ^ Term.string_of_vname xi)
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   454
  | reject_schematic (Abs (_, _, t)) = reject_schematic t
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   455
  | reject_schematic (t $ u) = (reject_schematic t; reject_schematic u)
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   456
  | reject_schematic _ = ();
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   457
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   458
fun expand_binds ctxt schematic =
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   459
  Variable.expand_binds ctxt #> (if schematic then I else tap reject_schematic);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   460
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   461
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   462
(* schematic type variables *)
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   463
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   464
val reject_tvars = (Term.map_types o Term.map_atyps)
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   465
  (fn TVar (xi, _) => error ("Illegal schematic type variable: " ^ Term.string_of_vname xi)
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   466
    | T => T);
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   467
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   468
6550
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   469
(* dummy patterns *)
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   470
18310
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   471
val prepare_dummies =
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   472
  let val next = ref 1 in
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   473
    fn t =>
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   474
      let val (i, u) = Term.replace_dummy_patterns (! next, t)
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   475
      in next := i; u end
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   476
  end;
6762
a9a515a43ae0 read_term/prop_pat: do not freeze;
wenzelm
parents: 6721
diff changeset
   477
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   478
fun reject_dummies t = Term.no_dummy_patterns t
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   479
  handle TERM _ => error "Illegal dummy pattern(s) in term";
6550
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   480
68f950b1a664 dummy patterns;
wenzelm
parents: 6528
diff changeset
   481
22763
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   482
(* decoding raw terms (syntax trees) *)
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   483
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   484
fun legacy_intern_skolem ctxt internal def_type x =    (* FIXME cleanup this mess *)
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   485
  let
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   486
    val sko = lookup_skolem ctxt x;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   487
    val is_const = can (Consts.read_const (consts_of ctxt)) x;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   488
    val is_free = is_some sko orelse not is_const;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   489
    val is_internal = internal x;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   490
    val is_declared = is_some (def_type (x, ~1));
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   491
    val res =
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   492
      if is_free andalso not is_internal then (no_skolem false x; sko)
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   493
      else ((no_skolem false x; ()) handle ERROR msg => warning msg;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   494
        if is_internal andalso is_declared then SOME x else NONE);
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   495
  in if is_some res then res else if is_declared then SOME x else NONE end;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   496
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   497
fun decode_term ctxt =
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   498
  let
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   499
    val thy = theory_of ctxt;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   500
    val get_sort = Sign.get_sort thy (Variable.def_sort ctxt);
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   501
    val map_const = try (#1 o Term.dest_Const o Consts.read_const (consts_of ctxt));
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   502
    val map_free = legacy_intern_skolem ctxt (K false) (Variable.def_type ctxt false);
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   503
    val map_type = Sign.intern_tycons thy;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   504
    val map_sort = Sign.intern_sort thy;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   505
  in Syntax.decode_term get_sort map_const map_free map_type map_sort end;
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   506
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   507
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   508
(* read terms *)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   509
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   510
local
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   511
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   512
fun append_env e1 e2 x = (case e2 x of NONE => e1 x | some => some);
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   513
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   514
fun gen_read' read app pattern schematic
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   515
    ctxt internal more_types more_sorts more_used s =
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13629
diff changeset
   516
  let
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   517
    val types = append_env (Variable.def_type ctxt pattern) more_types;
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   518
    val sorts = append_env (Variable.def_sort ctxt) more_sorts;
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
   519
    val used = fold Name.declare more_used (Variable.names_of ctxt);
14174
f3cafd2929d5 Methods rule_tac etc support static (Isar) contexts.
ballarin
parents: 13629
diff changeset
   520
  in
22763
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   521
    (read (pp ctxt) (syn_of ctxt) ctxt (types, sorts, used)
5c1752279f25 added decode_term (belongs to Syntax module);
wenzelm
parents: 22728
diff changeset
   522
        (legacy_intern_skolem ctxt internal types) s
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   523
      handle TERM (msg, _) => error msg)
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   524
    |> app (certify_consts ctxt)
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   525
    |> app (if pattern then I else expand_binds ctxt schematic)
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   526
    |> app (if pattern orelse schematic then I else reject_tvars)
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   527
    |> app (if pattern then prepare_dummies else reject_dummies)
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   528
  end;
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   529
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   530
fun gen_read read app pattern schematic ctxt =
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   531
  gen_read' read app pattern schematic ctxt (K false) (K NONE) (K NONE) [];
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   532
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   533
in
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   534
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   535
val read_termTs           = gen_read' (read_def_termTs false) (apfst o map) false false;
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   536
val read_termTs_schematic = gen_read' (read_def_termTs false) (apfst o map) false true;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   537
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   538
fun read_term_pats T ctxt =
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   539
  #1 o gen_read (read_def_termTs false) (apfst o map) true false ctxt o map (rpair T);
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   540
val read_prop_pats = read_term_pats propT;
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   541
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   542
fun read_prop_legacy ctxt =
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   543
  gen_read' (read_prop_thy true) I false false ctxt (K true) (K NONE) (K NONE) [];
14720
ceff6d4fb836 cleanup up read functions, include liberal versions;
wenzelm
parents: 14707
diff changeset
   544
19143
a64fef2d7073 put_thms: do_index;
wenzelm
parents: 19079
diff changeset
   545
val read_term            = gen_read (read_term_thy true) I false false;
a64fef2d7073 put_thms: do_index;
wenzelm
parents: 19079
diff changeset
   546
val read_prop            = gen_read (read_prop_thy true) I false false;
a64fef2d7073 put_thms: do_index;
wenzelm
parents: 19079
diff changeset
   547
val read_prop_schematic  = gen_read (read_prop_thy true) I false true;
a64fef2d7073 put_thms: do_index;
wenzelm
parents: 19079
diff changeset
   548
val read_terms           = gen_read (read_terms_thy true) map false false;
a64fef2d7073 put_thms: do_index;
wenzelm
parents: 19079
diff changeset
   549
fun read_props schematic = gen_read (read_props_thy true) map false schematic;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   550
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   551
end;
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   552
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   553
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   554
(* certify terms *)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   555
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   556
local
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   557
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   558
fun gen_cert prop pattern schematic ctxt t = t
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   559
  |> certify_consts ctxt
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   560
  |> (if pattern then I else expand_binds ctxt schematic)
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   561
  |> (fn t' => #1 (Sign.certify' false prop (pp ctxt) (consts_of ctxt) (theory_of ctxt) t')
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   562
    handle TYPE (msg, _, _) => error msg
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   563
      | TERM (msg, _) => error msg);
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   564
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   565
in
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   566
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   567
val cert_term = gen_cert false false false;
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   568
val cert_prop = gen_cert true false false;
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   569
val cert_props = map oo gen_cert true false;
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   570
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   571
fun cert_term_pats _ = map o gen_cert false true false;
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   572
val cert_prop_pats = map o gen_cert true true false;
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   573
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   574
end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   575
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   576
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   577
(* type inference *)
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   578
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   579
local
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   580
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
   581
fun gen_infer_types pattern ctxt ts =
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   582
  TypeInfer.infer_types (pp ctxt) (tsig_of ctxt) (try (Consts.the_constraint (consts_of ctxt)))
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
   583
    (Variable.def_type ctxt pattern) (Variable.names_of ctxt) (not pattern) (map (rpair dummyT) ts)
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
   584
  |> #1 |> map #1;
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   585
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   586
in
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   587
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   588
val infer_types = gen_infer_types false;
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
   589
val infer_types_pats = gen_infer_types true;
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   590
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   591
end;
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   592
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   593
18770
434f660d3068 renamed inferred_type to inferred_param;
wenzelm
parents: 18743
diff changeset
   594
(* inferred types of parameters *)
434f660d3068 renamed inferred_type to inferred_param;
wenzelm
parents: 18743
diff changeset
   595
434f660d3068 renamed inferred_type to inferred_param;
wenzelm
parents: 18743
diff changeset
   596
fun infer_type ctxt x =
22728
ecbbdf50df2f simplified ProofContext.infer_types(_pats);
wenzelm
parents: 22712
diff changeset
   597
  Term.fastype_of (singleton (infer_types ctxt) (Free (x, dummyT)));
18770
434f660d3068 renamed inferred_type to inferred_param;
wenzelm
parents: 18743
diff changeset
   598
434f660d3068 renamed inferred_type to inferred_param;
wenzelm
parents: 18743
diff changeset
   599
fun inferred_param x ctxt =
18619
fa61df848dea added infer_type, declared_type;
wenzelm
parents: 18609
diff changeset
   600
  let val T = infer_type ctxt x
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
   601
  in ((x, T), ctxt |> Variable.declare_term (Free (x, T))) end;
18619
fa61df848dea added infer_type, declared_type;
wenzelm
parents: 18609
diff changeset
   602
18770
434f660d3068 renamed inferred_type to inferred_param;
wenzelm
parents: 18743
diff changeset
   603
fun inferred_fixes ctxt =
20086
94ca946fb689 adapted to more efficient Name/Variable implementation;
wenzelm
parents: 20049
diff changeset
   604
  fold_map inferred_param (rev (map #2 (Variable.fixes_of ctxt))) ctxt;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   605
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   606
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   607
(* type and constant names *)
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   608
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   609
fun read_tyname ctxt c =
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
   610
  if Syntax.is_tid c then
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   611
    TFree (c, the_default (Sign.defaultS (theory_of ctxt)) (Variable.def_sort ctxt (c, ~1)))
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16348
diff changeset
   612
  else Sign.read_tyname (theory_of ctxt) c;
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   613
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   614
fun read_const ctxt c =
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   615
  (case lookup_skolem ctxt c of
21208
62ccdaf9369a replaced const_syntax by notation, which operates on terms;
wenzelm
parents: 21183
diff changeset
   616
    SOME x => Free (x, infer_type ctxt x)
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
   617
  | NONE => Consts.read_const (consts_of ctxt) c);
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   618
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
   619
9553
c2e3773475b6 norm_hhf results;
wenzelm
parents: 9540
diff changeset
   620
21610
52c0d3280798 removed obsolete (export_)standard;
wenzelm
parents: 21600
diff changeset
   621
(** export results **)
21531
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   622
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   623
fun common_export is_goal inner outer =
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   624
  map (Assumption.export is_goal inner outer) #>
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   625
  Variable.export inner outer;
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   626
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   627
val goal_export = common_export true;
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
   628
val export = common_export false;
12704
7bffaadc581e export_single;
wenzelm
parents: 12576
diff changeset
   629
21531
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   630
fun export_morphism inner outer =
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   631
  Assumption.export_morphism inner outer $>
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   632
  Variable.export_morphism inner outer;
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   633
43aa65a8a870 added export_(standard_)morphism;
wenzelm
parents: 21443
diff changeset
   634
15758
07e382399a96 binds/thms: do not store options, but delete from table;
wenzelm
parents: 15750
diff changeset
   635
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   636
(** bindings **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   637
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   638
(* simult_matches *)
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   639
19867
474cc9b49239 added declare_typ;
wenzelm
parents: 19847
diff changeset
   640
fun simult_matches ctxt (t, pats) =
474cc9b49239 added declare_typ;
wenzelm
parents: 19847
diff changeset
   641
  (case Seq.pull (Unify.matchers (theory_of ctxt) (map (rpair t) pats)) of
474cc9b49239 added declare_typ;
wenzelm
parents: 19847
diff changeset
   642
    NONE => error "Pattern match failed!"
474cc9b49239 added declare_typ;
wenzelm
parents: 19847
diff changeset
   643
  | SOME (env, _) => map (apsnd snd) (Envir.alist_of env));
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   644
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   645
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   646
(* add_binds(_i) *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   647
7925
8c50b68b890b warn_extra_tfrees (after declare_term);
wenzelm
parents: 7679
diff changeset
   648
local
8c50b68b890b warn_extra_tfrees (after declare_term);
wenzelm
parents: 7679
diff changeset
   649
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   650
fun gen_bind prep (xi as (x, _), raw_t) ctxt =
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   651
  ctxt |> Variable.add_binds [(xi, Option.map (prep ctxt) raw_t)];
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   652
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
   653
in
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
   654
20330
wenzelm
parents: 20310
diff changeset
   655
fun drop_schematic (b as (xi, SOME t)) = if Term.exists_subterm is_Var t then (xi, NONE) else b
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   656
  | drop_schematic b = b;
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   657
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   658
val add_binds = fold (gen_bind read_term);
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   659
val add_binds_i = fold (gen_bind cert_term);
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   660
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16348
diff changeset
   661
fun auto_bind f ts ctxt = ctxt |> add_binds_i (map drop_schematic (f (theory_of ctxt) ts));
12147
64e69a8a945f adapted auto_bind_goal, auto_bind_facts;
wenzelm
parents: 12130
diff changeset
   662
val auto_bind_goal = auto_bind AutoBind.goal;
64e69a8a945f adapted auto_bind_goal, auto_bind_facts;
wenzelm
parents: 12130
diff changeset
   663
val auto_bind_facts = auto_bind AutoBind.facts;
7925
8c50b68b890b warn_extra_tfrees (after declare_term);
wenzelm
parents: 7679
diff changeset
   664
8c50b68b890b warn_extra_tfrees (after declare_term);
wenzelm
parents: 7679
diff changeset
   665
end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   666
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   667
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   668
(* match_bind(_i) *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   669
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   670
local
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   671
17860
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   672
fun prep_bind prep_pats (raw_pats, t) ctxt =
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   673
  let
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   674
    val ctxt' = Variable.declare_term t ctxt;
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   675
    val pats = prep_pats (Term.fastype_of t) ctxt' raw_pats;
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   676
    val binds = simult_matches ctxt' (t, pats);
17860
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   677
  in (binds, ctxt') end;
7670
e302e4269087 added cert_skolem;
wenzelm
parents: 7663
diff changeset
   678
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   679
fun gen_binds prep_terms prep_pats gen raw_binds ctxt =
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   680
  let
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   681
    val ts = prep_terms ctxt (map snd raw_binds);
17860
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   682
    val (binds, ctxt') =
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19422
diff changeset
   683
      apfst flat (fold_map (prep_bind prep_pats) (map fst raw_binds ~~ ts) ctxt);
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   684
    val binds' =
19916
3bbb9cc5d4f1 export: simultaneous facts, refer to Variable.export;
wenzelm
parents: 19897
diff changeset
   685
      if gen then map #1 binds ~~ Variable.exportT_terms ctxt' ctxt (map #2 binds)
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   686
      else binds;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   687
    val binds'' = map (apsnd SOME) binds';
18310
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   688
    val ctxt'' =
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   689
      tap (Variable.warn_extra_tfrees ctxt)
18310
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   690
       (if gen then
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   691
          ctxt (*sic!*) |> fold Variable.declare_term (map #2 binds') |> add_binds_i binds''
18310
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   692
        else ctxt' |> add_binds_i binds'');
b00c9120f6bd match_bind(_i): return terms;
wenzelm
parents: 18255
diff changeset
   693
  in (ts, ctxt'') end;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   694
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   695
in
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   696
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   697
val match_bind = gen_binds read_terms read_term_pats;
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   698
val match_bind_i = gen_binds (map o cert_term) cert_term_pats;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   699
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   700
end;
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   701
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   702
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   703
(* propositions with patterns *)
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   704
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   705
local
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   706
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   707
fun prep_propp schematic prep_props prep_pats (context, args) =
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   708
  let
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   709
    fun prep (_, raw_pats) (ctxt, prop :: props) =
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   710
          let val ctxt' = Variable.declare_term prop ctxt
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   711
          in ((prop, prep_pats ctxt' raw_pats), (ctxt', props)) end
17860
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   712
      | prep _ _ = sys_error "prep_propp";
b4cf247ea0d2 note_thmss, read/cert_vars etc.: natural argument order;
wenzelm
parents: 17756
diff changeset
   713
    val (propp, (context', _)) = (fold_map o fold_map) prep args
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19422
diff changeset
   714
      (context, prep_props schematic context (maps (map fst) args));
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   715
  in (context', propp) end;
5935
6a82c8a1808f term_pat vs. prop_pat;
wenzelm
parents: 5919
diff changeset
   716
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   717
fun gen_bind_propp prepp (ctxt, raw_args) =
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   718
  let
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   719
    val (ctxt', args) = prepp (ctxt, raw_args);
19585
70a1ce3b23ae removed 'concl is' patterns;
wenzelm
parents: 19543
diff changeset
   720
    val binds = flat (flat (map (map (simult_matches ctxt')) args));
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   721
    val propss = map (map #1) args;
8616
90d2fed59be1 support Hindley-Milner polymorphisms in binds and facts;
wenzelm
parents: 8462
diff changeset
   722
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   723
    (*generalize result: context evaluated now, binds added later*)
19916
3bbb9cc5d4f1 export: simultaneous facts, refer to Variable.export;
wenzelm
parents: 19897
diff changeset
   724
    val gen = Variable.exportT_terms ctxt' ctxt;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   725
    fun gen_binds c = c |> add_binds_i (map #1 binds ~~ map SOME (gen (map #2 binds)));
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
   726
  in (ctxt' |> add_binds_i (map (apsnd SOME) binds), (propss, gen_binds)) end;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   727
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   728
in
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   729
11925
4747b4b84093 added read_prop_schematic;
wenzelm
parents: 11918
diff changeset
   730
val read_propp = prep_propp false read_props read_prop_pats;
4747b4b84093 added read_prop_schematic;
wenzelm
parents: 11918
diff changeset
   731
val cert_propp = prep_propp false cert_props cert_prop_pats;
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   732
val read_propp_schematic = prep_propp true read_props read_prop_pats;
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   733
val cert_propp_schematic = prep_propp true cert_props cert_prop_pats;
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   734
11925
4747b4b84093 added read_prop_schematic;
wenzelm
parents: 11918
diff changeset
   735
val bind_propp = gen_bind_propp read_propp;
4747b4b84093 added read_prop_schematic;
wenzelm
parents: 11918
diff changeset
   736
val bind_propp_i = gen_bind_propp cert_propp;
4747b4b84093 added read_prop_schematic;
wenzelm
parents: 11918
diff changeset
   737
val bind_propp_schematic = gen_bind_propp read_propp_schematic;
10554
81edb1d201ab schematic props;
wenzelm
parents: 10465
diff changeset
   738
val bind_propp_schematic_i = gen_bind_propp cert_propp_schematic;
6789
0e5a965de17a auto_bind_goal, auto_bind_facts;
wenzelm
parents: 6762
diff changeset
   739
10465
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   740
end;
4aa6f8b5cdc4 added read_terms, read_props (simulataneous type-inference);
wenzelm
parents: 10381
diff changeset
   741
6789
0e5a965de17a auto_bind_goal, auto_bind_facts;
wenzelm
parents: 6762
diff changeset
   742
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   743
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   744
(** theorems **)
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   745
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   746
(* fact_tac *)
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   747
18122
d5a876195499 removed export_plain;
wenzelm
parents: 18043
diff changeset
   748
fun comp_incr_tac [] _ st = no_tac st
d5a876195499 removed export_plain;
wenzelm
parents: 18043
diff changeset
   749
  | comp_incr_tac (th :: ths) i st =
d5a876195499 removed export_plain;
wenzelm
parents: 18043
diff changeset
   750
      (Goal.compose_hhf_tac (Drule.incr_indexes st th) i APPEND comp_incr_tac ths i) st;
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   751
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21681
diff changeset
   752
fun fact_tac facts = Goal.norm_hhf_tac THEN' comp_incr_tac facts;
18122
d5a876195499 removed export_plain;
wenzelm
parents: 18043
diff changeset
   753
d5a876195499 removed export_plain;
wenzelm
parents: 18043
diff changeset
   754
fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) =>
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   755
  let
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   756
    val index = fact_index_of ctxt;
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   757
    val facts = FactIndex.could_unify index (Term.strip_all_body goal);
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   758
  in fact_tac facts i end);
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   759
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   760
6091
e3cdbd929a24 eliminated tthm type and Attribute structure;
wenzelm
parents: 5994
diff changeset
   761
(* get_thm(s) *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   762
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   763
fun retrieve_thms _ pick ctxt (Fact s) =
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   764
      let
20049
f48c4a3a34bc Goal.prove: context;
wenzelm
parents: 20012
diff changeset
   765
        val th = Goal.prove ctxt [] [] (read_prop ctxt s) (K (ALLGOALS (some_fact_tac ctxt)))
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   766
          handle ERROR msg => cat_error msg "Failed to retrieve literal fact.";
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   767
      in pick "" [th] end
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   768
  | retrieve_thms from_thy pick ctxt xthmref =
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   769
      let
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   770
        val thy = theory_of ctxt;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   771
        val (space, tab) = #1 (thms_of ctxt);
16501
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   772
        val thmref = PureThy.map_name_of_thmref (NameSpace.intern space) xthmref;
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   773
        val name = PureThy.name_of_thmref thmref;
fec0cf020bad thmref: Name vs. NameSelection;
wenzelm
parents: 16458
diff changeset
   774
      in
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17384
diff changeset
   775
        (case Symtab.lookup tab name of
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   776
          SOME ths => map (Thm.transfer thy) (PureThy.select_thm thmref ths)
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
   777
        | NONE => from_thy thy xthmref) |> pick name
18042
f9890c615c0e added fact_tac, some_fact_tac;
wenzelm
parents: 17977
diff changeset
   778
      end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   779
9566
0874bf3a909d thms closure;
wenzelm
parents: 9553
diff changeset
   780
val get_thm = retrieve_thms PureThy.get_thms PureThy.single_thm;
0874bf3a909d thms closure;
wenzelm
parents: 9553
diff changeset
   781
val get_thm_closure = retrieve_thms PureThy.get_thms_closure PureThy.single_thm;
0874bf3a909d thms closure;
wenzelm
parents: 9553
diff changeset
   782
val get_thms = retrieve_thms PureThy.get_thms (K I);
0874bf3a909d thms closure;
wenzelm
parents: 9553
diff changeset
   783
val get_thms_closure = retrieve_thms PureThy.get_thms_closure (K I);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   784
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   785
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   786
(* valid_thms *)
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   787
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   788
fun valid_thms ctxt (name, ths) =
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   789
  (case try (fn () => get_thms ctxt (Name name)) () of
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   790
    NONE => false
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
   791
  | SOME ths' => Thm.eq_thms (ths, ths'));
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   792
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   793
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   794
(* lthms_containing *)
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   795
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   796
fun lthms_containing ctxt spec =
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   797
  FactIndex.find (fact_index_of ctxt) spec
21569
a0d0ea84521d simplified '?' operator;
wenzelm
parents: 21531
diff changeset
   798
  |> map (fn (name, ths) =>
a0d0ea84521d simplified '?' operator;
wenzelm
parents: 21531
diff changeset
   799
    if valid_thms ctxt (name, ths) then (name, ths)
a0d0ea84521d simplified '?' operator;
wenzelm
parents: 21531
diff changeset
   800
    else (NameSpace.hidden (if name = "" then "unnamed" else name), ths));
16031
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   801
fbf3471214d6 moved everything related to thms_containing to find_theorems.ML;
wenzelm
parents: 15979
diff changeset
   802
13425
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
   803
(* name space operations *)
12309
03e9287be350 name space for local thms (export cond_extern, qualified);
wenzelm
parents: 12291
diff changeset
   804
22352
f15118a79c0e add_path for naming in proof contexts
haftmann
parents: 21824
diff changeset
   805
val add_path        = map_naming o NameSpace.add_path;
19062
0fd52e819c24 replaced qualified_force_prefix to sticky_prefix;
wenzelm
parents: 19033
diff changeset
   806
val no_base_names   = map_naming NameSpace.no_base_names;
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
   807
val qualified_names = map_naming NameSpace.qualified_names;
19062
0fd52e819c24 replaced qualified_force_prefix to sticky_prefix;
wenzelm
parents: 19033
diff changeset
   808
val sticky_prefix   = map_naming o NameSpace.sticky_prefix;
0fd52e819c24 replaced qualified_force_prefix to sticky_prefix;
wenzelm
parents: 19033
diff changeset
   809
val restore_naming  = map_naming o K o naming_of;
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   810
val reset_naming    = map_naming (K local_naming);
12309
03e9287be350 name space for local thms (export cond_extern, qualified);
wenzelm
parents: 12291
diff changeset
   811
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   812
fun hide_thms fully names = map_thms (fn ((space, tab), index) =>
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   813
  ((fold (NameSpace.hide fully) names space, tab), index));
13425
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
   814
12309
03e9287be350 name space for local thms (export cond_extern, qualified);
wenzelm
parents: 12291
diff changeset
   815
17360
fa1f262dbc4e added add_view, export_view (supercedes adhoc view arguments);
wenzelm
parents: 17221
diff changeset
   816
(* put_thms *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   817
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   818
fun update_thms _ ("", NONE) ctxt = ctxt
21648
c8a0370c9b93 more careful indexing of local facts;
wenzelm
parents: 21643
diff changeset
   819
  | update_thms opts ("", SOME ths) ctxt = ctxt |> map_thms (fn (facts, index) =>
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   820
      let
21648
c8a0370c9b93 more careful indexing of local facts;
wenzelm
parents: 21643
diff changeset
   821
        val index' = uncurry FactIndex.add_local opts ("", ths) index;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   822
      in (facts, index') end)
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   823
  | update_thms _ (bname, NONE) ctxt = ctxt |> map_thms (fn ((space, tab), index) =>
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   824
      let
19387
6af442fa80c3 added full_name;
wenzelm
parents: 19371
diff changeset
   825
        val name = full_name ctxt bname;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   826
        val tab' = Symtab.delete_safe name tab;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   827
      in ((space, tab'), index) end)
21648
c8a0370c9b93 more careful indexing of local facts;
wenzelm
parents: 21643
diff changeset
   828
  | update_thms opts (bname, SOME ths) ctxt = ctxt |> map_thms (fn ((space, tab), index) =>
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   829
      let
19387
6af442fa80c3 added full_name;
wenzelm
parents: 19371
diff changeset
   830
        val name = full_name ctxt bname;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   831
        val space' = NameSpace.declare (naming_of ctxt) name space;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   832
        val tab' = Symtab.update (name, ths) tab;
21648
c8a0370c9b93 more careful indexing of local facts;
wenzelm
parents: 21643
diff changeset
   833
        val index' = uncurry FactIndex.add_local opts (name, ths) index;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   834
      in ((space', tab'), index') end);
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   835
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   836
fun put_thms thms ctxt =
21648
c8a0370c9b93 more careful indexing of local facts;
wenzelm
parents: 21643
diff changeset
   837
  ctxt |> map_naming (K local_naming) |> update_thms (true, false) thms |> restore_naming ctxt;
19079
9a7678a0736d added put_thms_internal: local_naming, no fact index;
wenzelm
parents: 19062
diff changeset
   838
7606
7905a74eb068 added reset_thms;
wenzelm
parents: 7575
diff changeset
   839
14564
3667b4616e9a renamed have_thms to note_thms;
wenzelm
parents: 14508
diff changeset
   840
(* note_thmss *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   841
12711
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
   842
local
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
   843
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   844
fun gen_note_thmss get k = fold_map (fn ((bname, more_attrs), raw_facts) => fn ctxt =>
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   845
  let
21643
bdf3e74727df note_thmss: added kind tag and non-official name;
wenzelm
parents: 21622
diff changeset
   846
    val name = full_name ctxt bname;
bdf3e74727df note_thmss: added kind tag and non-official name;
wenzelm
parents: 21622
diff changeset
   847
    val facts = PureThy.name_thmss false name (map (apfst (get ctxt)) raw_facts);
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   848
    fun app (th, attrs) x =
21643
bdf3e74727df note_thmss: added kind tag and non-official name;
wenzelm
parents: 21622
diff changeset
   849
      swap (foldl_map (Thm.proof_attributes (attrs @ more_attrs @ [PureThy.kind k])) (x, th));
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
   850
    val (res, ctxt') = fold_map app facts ctxt;
21643
bdf3e74727df note_thmss: added kind tag and non-official name;
wenzelm
parents: 21622
diff changeset
   851
    val thms = PureThy.name_thms false false name (flat res);
21648
c8a0370c9b93 more careful indexing of local facts;
wenzelm
parents: 21643
diff changeset
   852
  in ((bname, thms), ctxt' |> update_thms (is_stmt ctxt, true) (bname, SOME thms)) end);
12711
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
   853
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
   854
in
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
   855
21622
5825a59785f4 made SML/NJ happy
haftmann
parents: 21612
diff changeset
   856
fun note_thmss k = gen_note_thmss get_thms k;
5825a59785f4 made SML/NJ happy
haftmann
parents: 21612
diff changeset
   857
fun note_thmss_i k = gen_note_thmss (K I) k;
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15624
diff changeset
   858
12711
6a9412dd7d24 have_thmss vs. have_thmss_i;
wenzelm
parents: 12704
diff changeset
   859
end;
9196
1f6f66fe777a facts: handle multiple lists of arguments;
wenzelm
parents: 9133
diff changeset
   860
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   861
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   862
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   863
(** parameters **)
17360
fa1f262dbc4e added add_view, export_view (supercedes adhoc view arguments);
wenzelm
parents: 17221
diff changeset
   864
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   865
(* variables *)
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   866
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   867
fun declare_var (x, opt_T, mx) ctxt =
22701
4346f230283d proper interface infer_types(_pat);
wenzelm
parents: 22678
diff changeset
   868
  let val T = (case opt_T of SOME T => T | NONE => Syntax.mixfixT mx)
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
   869
  in ((x, T, mx), ctxt |> Variable.declare_constraints (Free (x, T))) end;
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   870
10381
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
   871
local
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
   872
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   873
fun prep_vars prep_typ internal legacy =
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   874
  fold_map (fn (raw_x, raw_T, raw_mx) => fn ctxt =>
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   875
    let
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   876
      val (x, mx) = Syntax.const_mixfix raw_x raw_mx;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   877
      val _ =
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   878
        if not legacy andalso not (Syntax.is_identifier (no_skolem internal x)) then
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   879
          error ("Illegal variable name: " ^ quote x)
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   880
        else ();
12504
5b46173df7ad export used_types;
wenzelm
parents: 12414
diff changeset
   881
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   882
      fun cond_tvars T =
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   883
        if internal then T
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
   884
        else Type.no_tvars T handle TYPE (msg, _, _) => error msg;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   885
      val opt_T = Option.map (cond_tvars o prep_typ ctxt) raw_T;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   886
      val var = (x, opt_T, mx);
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   887
    in (var, ctxt |> declare_var var |> #2) end);
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   888
10381
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
   889
in
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
   890
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   891
val read_vars        = prep_vars read_typ false false;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   892
val cert_vars        = prep_vars cert_typ true false;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   893
val read_vars_legacy = prep_vars read_typ true true;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   894
val cert_vars_legacy = prep_vars cert_typ true true;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   895
10381
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
   896
end;
4dfbcad19bfb fixed two obscurities of "fix": predeclare_terms;
wenzelm
parents: 9733
diff changeset
   897
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   898
19681
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   899
(* authentic constants *)
19663
459848022d6e export consts_of;
wenzelm
parents: 19585
diff changeset
   900
21208
62ccdaf9369a replaced const_syntax by notation, which operates on terms;
wenzelm
parents: 21183
diff changeset
   901
fun const_syntax ctxt (Free (x, T), mx) = SOME (true, (x, T, mx))
21667
ce813b82c88b add_notation: permissive about undeclared consts;
wenzelm
parents: 21648
diff changeset
   902
  | const_syntax ctxt (Const (c, _), mx) =
ce813b82c88b add_notation: permissive about undeclared consts;
wenzelm
parents: 21648
diff changeset
   903
      Option.map (pair false) (try (Consts.syntax (consts_of ctxt)) (c, mx))
21208
62ccdaf9369a replaced const_syntax by notation, which operates on terms;
wenzelm
parents: 21183
diff changeset
   904
  | const_syntax _ _ = NONE;
62ccdaf9369a replaced const_syntax by notation, which operates on terms;
wenzelm
parents: 21183
diff changeset
   905
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21752
diff changeset
   906
fun context_const_ast_tr ctxt [Syntax.Variable c] =
19681
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   907
      let
21772
7c7ade4f537b advanced translation functions: Proof.context;
wenzelm
parents: 21752
diff changeset
   908
        val consts = consts_of ctxt;
19681
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   909
        val c' = Consts.intern consts c;
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   910
        val _ = Consts.the_constraint consts c' handle TYPE (msg, _, _) => error msg;
19681
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   911
      in Syntax.Constant (Syntax.constN ^ c') end
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   912
  | context_const_ast_tr _ asts = raise Syntax.AST ("const_ast_tr", asts);
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   913
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   914
val _ = Context.add_setup
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   915
 (Sign.add_syntax
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   916
   [("_context_const", "id => 'a", Delimfix "CONST _"),
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   917
    ("_context_const", "longid => 'a", Delimfix "CONST _")] #>
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   918
  Sign.add_advanced_trfuns ([("_context_const", context_const_ast_tr)], [], [], []));
871167b2c70e added CONST syntax;
wenzelm
parents: 19673
diff changeset
   919
19663
459848022d6e export consts_of;
wenzelm
parents: 19585
diff changeset
   920
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   921
(* notation *)
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   922
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   923
fun add_notation mode args ctxt =
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   924
  ctxt |> map_syntax
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   925
    (LocalSyntax.add_modesyntax (theory_of ctxt) mode (map_filter (const_syntax ctxt) args));
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   926
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   927
fun target_notation mode args phi =
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   928
  let val args' = filter (fn (t, _) => Term.equiv_types (t, Morphism.term phi t)) args;
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   929
  in Context.mapping (Sign.add_notation mode args') (add_notation mode args') end;
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   930
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   931
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   932
(* abbreviations *)
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   933
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   934
val set_expand_abbrevs = map_consts o apsnd o Consts.set_expand;
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   935
fun read_term_abbrev ctxt = read_term (set_expand_abbrevs false ctxt);
19371
32fc9743803a add_abbrevs(_i): support print mode;
wenzelm
parents: 19310
diff changeset
   936
21704
f4fe6e5a3ee6 simplified add_abbrev -- single argument;
wenzelm
parents: 21696
diff changeset
   937
fun add_abbrev mode (c, raw_t) ctxt =
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   938
  let
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   939
    val t0 = cert_term (ctxt |> set_expand_abbrevs false) raw_t
21681
8b8edcdb4da8 add_abbrevs: moved common parts to consts.ML;
wenzelm
parents: 21667
diff changeset
   940
      handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote c);
20008
8d9d770e1f06 add_abbrevs/polymorphic: Variable.exportT_terms avoids over-generalization;
wenzelm
parents: 19916
diff changeset
   941
    val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0];
21807
a59f083632a7 add_abbrev: removed Assumption.add_assms (danger of inconsistent naming);
wenzelm
parents: 21803
diff changeset
   942
    val ((lhs, rhs), consts') = consts_of ctxt
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
   943
      |> Consts.abbreviate (pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode (c, t);
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   944
  in
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   945
    ctxt
21681
8b8edcdb4da8 add_abbrevs: moved common parts to consts.ML;
wenzelm
parents: 21667
diff changeset
   946
    |> map_consts (apsnd (K consts'))
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   947
    |> Variable.declare_term rhs
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   948
    |> pair (lhs, rhs)
21704
f4fe6e5a3ee6 simplified add_abbrev -- single argument;
wenzelm
parents: 21696
diff changeset
   949
  end;
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   950
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   951
fun target_abbrev prmode ((c, mx), rhs) phi =
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   952
  let
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   953
    val c' = Morphism.name phi c;
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   954
    val rhs' = Morphism.term phi rhs;
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   955
    val arg' = (c', rhs');
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   956
  in
21824
153fad1e7318 target_abbrev: internal mode for abbrevs;
wenzelm
parents: 21814
diff changeset
   957
    Context.mapping_result
153fad1e7318 target_abbrev: internal mode for abbrevs;
wenzelm
parents: 21814
diff changeset
   958
      (Sign.add_abbrev Syntax.internalM arg') (add_abbrev Syntax.internalM arg')
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   959
    #-> (fn (lhs, _) =>
22670
c803b2696ada added Morphism.transform/form (generic non-sense);
wenzelm
parents: 22587
diff changeset
   960
      Term.equiv_types (rhs, rhs') ? Morphism.form (target_notation prmode [(lhs, mx)]))
21744
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   961
  end;
b790ce4c21c2 added target_notation/abbrev;
wenzelm
parents: 21728
diff changeset
   962
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   963
fun local_abbrev (x, rhs) =
21814
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   964
  Variable.add_fixes [x] #-> (fn [x'] =>
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   965
    let
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   966
      val T = Term.fastype_of rhs;
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   967
      val lhs = Free (x', T);
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   968
    in
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   969
      Variable.declare_term lhs #>
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   970
      Variable.declare_term rhs #>
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   971
      Assumption.add_assms (K (K (I, Envir.expand_term_frees [((x', T), rhs)]))) [] #> snd #>
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   972
      pair (lhs, rhs)
d7bb902beb07 local_abbrev: proper fix/declare of local entities;
wenzelm
parents: 21807
diff changeset
   973
    end);
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
   974
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   975
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   976
(* fixes *)
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   977
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   978
local
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   979
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   980
fun prep_mixfix (x, T, mx) =
19019
412009243085 added map_theory;
wenzelm
parents: 19001
diff changeset
   981
  if mx <> NoSyn andalso mx <> Structure andalso
20086
94ca946fb689 adapted to more efficient Name/Variable implementation;
wenzelm
parents: 20049
diff changeset
   982
      (can Name.dest_internal x orelse can Name.dest_skolem x) then
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   983
    error ("Illegal mixfix syntax for internal/skolem constant " ^ quote x)
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   984
  else (true, (x, T, mx));
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   985
18844
9dd789841018 invent_fixes: merely enter body temporarily;
wenzelm
parents: 18827
diff changeset
   986
fun gen_fixes prep raw_vars ctxt =
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   987
  let
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   988
    val (vars, _) = prep raw_vars ctxt;
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   989
    val (xs', ctxt') = Variable.add_fixes (map #1 vars) ctxt;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   990
  in
22712
8f2ba236918b replaced read_term_legacy by read_prop_legacy;
wenzelm
parents: 22701
diff changeset
   991
    ctxt'
19001
64e4b5bc6443 tuned comment;
wenzelm
parents: 18971
diff changeset
   992
    |> fold_map declare_var (map2 (fn x' => fn (_, T, mx) => (x', T, mx)) xs' vars)
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
   993
    |-> (map_syntax o LocalSyntax.add_syntax (theory_of ctxt) o map prep_mixfix)
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
   994
    |> pair xs'
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   995
  end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
   996
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
   997
in
7679
99912beb8fa0 improved 'fix' / Skolem interfaces;
wenzelm
parents: 7670
diff changeset
   998
18844
9dd789841018 invent_fixes: merely enter body temporarily;
wenzelm
parents: 18827
diff changeset
   999
val add_fixes = gen_fixes read_vars;
9dd789841018 invent_fixes: merely enter body temporarily;
wenzelm
parents: 18827
diff changeset
  1000
val add_fixes_i = gen_fixes cert_vars;
9dd789841018 invent_fixes: merely enter body temporarily;
wenzelm
parents: 18827
diff changeset
  1001
val add_fixes_legacy = gen_fixes cert_vars_legacy;
8096
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1002
4da940d100f5 TypeInfer.logicT;
wenzelm
parents: 7928
diff changeset
  1003
end;
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1004
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1005
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1006
(* fixes vs. frees *)
12016
425289df84d3 fix_frees;
wenzelm
parents: 12008
diff changeset
  1007
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1008
fun auto_fixes (arg as (ctxt, (propss, x))) =
21370
d9dd7b4e5e69 replaced Variable.fix_frees by Variable.auto_fixes (depends on body mode);
wenzelm
parents: 21269
diff changeset
  1009
  ((fold o fold) Variable.auto_fixes propss ctxt, (propss, x));
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1010
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1011
fun bind_fixes xs ctxt =
9291
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1012
  let
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1013
    val (_, ctxt') = ctxt |> add_fixes_i (map (fn x => (x, NONE, NoSyn)) xs);
9291
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1014
    fun bind (t as Free (x, T)) =
18340
3fdff270aa04 removed fixed_tr: now handled in syntax module;
wenzelm
parents: 18310
diff changeset
  1015
          if member (op =) xs x then
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15456
diff changeset
  1016
            (case lookup_skolem ctxt' x of SOME x' => Free (x', T) | NONE => t)
9291
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1017
          else t
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1018
      | bind (t $ u) = bind t $ bind u
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1019
      | bind (Abs (x, T, t)) = Abs (x, T, bind t)
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1020
      | bind a = a;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1021
  in (bind, ctxt') end;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1022
9291
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1023
23705d14be8f "_i" arguments now expected to have skolems already internalized;
wenzelm
parents: 9274
diff changeset
  1024
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1025
(** assumptions **)
18187
ec44df8ffd21 added revert_skolem, mk_def, add_def;
wenzelm
parents: 18152
diff changeset
  1026
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1027
local
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1028
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1029
fun gen_assms prepp exp args ctxt =
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1030
  let
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1031
    val cert = Thm.cterm_of (theory_of ctxt);
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1032
    val (propss, ctxt1) = swap (prepp (ctxt, map snd args));
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1033
    val _ = Variable.warn_extra_tfrees ctxt ctxt1;
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1034
    val (premss, ctxt2) = fold_burrow (Assumption.add_assms exp o map cert) propss ctxt1;
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1035
  in
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1036
    ctxt2
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1037
    |> auto_bind_facts (flat propss)
21443
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
  1038
    |> put_thms (AutoBind.premsN, SOME (Assumption.prems_of ctxt2))
cc5095d57da4 added stmt mode, which affects naming/indexing of local facts;
wenzelm
parents: 21370
diff changeset
  1039
    |> note_thmss_i Thm.assumptionK (map fst args ~~ map (map (fn th => ([th], []))) premss)
20234
7e0693474bcd added legacy_pretty_thm (with fall-back on ProtoPure.thy);
wenzelm
parents: 20209
diff changeset
  1040
  end;
20209
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1041
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1042
in
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1043
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1044
val add_assms = gen_assms (apsnd #1 o bind_propp);
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1045
val add_assms_i = gen_assms (apsnd #1 o bind_propp_i);
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1046
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1047
end;
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1048
974d98969ba6 moved pprint functions to Isar/proof_display.ML;
wenzelm
parents: 20163
diff changeset
  1049
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1050
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1051
(** cases **)
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1052
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1053
local
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1054
16668
fdb4992cf1d2 avoid polyeq;
wenzelm
parents: 16540
diff changeset
  1055
fun rem_case name = remove (fn (x: string, (y, _)) => x = y) name;
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1056
18476
49dde7b7b14a cases: main is_proper flag;
wenzelm
parents: 18428
diff changeset
  1057
fun add_case _ ("", _) cases = cases
49dde7b7b14a cases: main is_proper flag;
wenzelm
parents: 18428
diff changeset
  1058
  | add_case _ (name, NONE) cases = rem_case name cases
49dde7b7b14a cases: main is_proper flag;
wenzelm
parents: 18428
diff changeset
  1059
  | add_case is_proper (name, SOME c) cases = (name, (c, is_proper)) :: rem_case name cases;
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1060
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
  1061
fun prep_case name fxs c =
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1062
  let
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1063
    fun replace (opt_x :: xs) ((y, T) :: ys) = (the_default y opt_x, T) :: replace xs ys
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1064
      | replace [] ys = ys
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
  1065
      | replace (_ :: _) [] = error ("Too many parameters for case " ^ quote name);
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1066
    val RuleCases.Case {fixes, assumes, binds, cases} = c;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1067
    val fixes' = replace fxs fixes;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1068
    val binds' = map drop_schematic binds;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1069
  in
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1070
    if null (fold (Term.add_tvarsT o snd) fixes []) andalso
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1071
      null (fold (fold Term.add_vars o snd) assumes []) then
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1072
        RuleCases.Case {fixes = fixes', assumes = assumes, binds = binds', cases = cases}
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
  1073
    else error ("Illegal schematic variable(s) in case " ^ quote name)
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1074
  end;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1075
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1076
fun fix (x, T) ctxt =
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1077
  let
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1078
    val (bind, ctxt') = bind_fixes [x] ctxt;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1079
    val t = bind (Free (x, T));
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
  1080
  in (t, ctxt' |> Variable.declare_constraints t) end;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1081
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1082
in
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1083
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1084
fun add_cases is_proper = map_cases o fold (add_case is_proper);
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1085
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1086
fun case_result c ctxt =
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1087
  let
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1088
    val RuleCases.Case {fixes, ...} = c;
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1089
    val (ts, ctxt') = ctxt |> fold_map fix fixes;
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1090
    val RuleCases.Case {assumes, binds, cases, ...} = RuleCases.apply ts c;
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1091
  in
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1092
    ctxt'
18699
f3bfe81b6e58 case_result: drop_schematic, i.e. be permissive about illegal binds;
wenzelm
parents: 18678
diff changeset
  1093
    |> add_binds_i (map drop_schematic binds)
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1094
    |> add_cases true (map (apsnd SOME) cases)
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1095
    |> pair (assumes, (binds, cases))
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1096
  end;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1097
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1098
val apply_case = apfst fst oo case_result;
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1099
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1100
fun get_case ctxt name xs =
17184
3d80209e9a53 use AList operations;
wenzelm
parents: 17072
diff changeset
  1101
  (case AList.lookup (op =) (cases_of ctxt) name of
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
  1102
    NONE => error ("Unknown case: " ^ quote name)
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18672
diff changeset
  1103
  | SOME (c, _) => prep_case name xs c);
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1104
16147
59177d5bcb6f renamed cond_extern to extern;
wenzelm
parents: 16031
diff changeset
  1105
end;
8373
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1106
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1107
e7237c8fe29e handling of local contexts: print_cases, get_case, add_cases;
wenzelm
parents: 8186
diff changeset
  1108
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1109
(** print context information **)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1110
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1111
val debug = ref false;
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1112
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1113
val verbose = ref false;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1114
fun verb f x = if ! verbose then f (x ()) else [];
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1115
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1116
fun setmp_verbose f x = Library.setmp verbose true f x;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1117
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1118
12072
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1119
(* local syntax *)
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1120
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
  1121
val print_syntax = Syntax.print_syntax o syn_of;
12072
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1122
4281198fb8cd local syntax: add_syntax, proper read/pretty functions;
wenzelm
parents: 12066
diff changeset
  1123
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1124
(* abbreviations *)
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1125
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1126
fun pretty_abbrevs show_globals ctxt =
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1127
  let
19033
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
  1128
    val ((_, globals), (space, consts)) =
24e251657e56 consts: maintain thy version for efficient transfer;
wenzelm
parents: 19019
diff changeset
  1129
      pairself (#constants o Consts.dest) (#consts (rep_context ctxt));
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
  1130
    fun add_abbr (_, (_, NONE)) = I
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
  1131
      | add_abbr (c, (T, SOME (t, _))) =
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1132
          if not show_globals andalso Symtab.defined globals c then I
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1133
          else cons (c, Logic.mk_equals (Const (c, T), t));
21803
bcef7eb50551 notation: Term.equiv_types;
wenzelm
parents: 21772
diff changeset
  1134
    val abbrevs = NameSpace.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1135
  in
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1136
    if null abbrevs andalso not (! verbose) then []
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1137
    else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1138
  end;
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1139
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1140
val print_abbrevs = Pretty.writeln o Pretty.chunks o pretty_abbrevs true;
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1141
18971
f95650f3b5bf added local consts component;
wenzelm
parents: 18953
diff changeset
  1142
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1143
(* term bindings *)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1144
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1145
fun pretty_binds ctxt =
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1146
  let
19897
fe661eb3b0e7 ProofContext: moved variable operations to struct Variable;
wenzelm
parents: 19882
diff changeset
  1147
    val binds = Variable.binds_of ctxt;
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1148
    fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t));
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1149
  in
15758
07e382399a96 binds/thms: do not store options, but delete from table;
wenzelm
parents: 15750
diff changeset
  1150
    if Vartab.is_empty binds andalso not (! verbose) then []
07e382399a96 binds/thms: do not store options, but delete from table;
wenzelm
parents: 15750
diff changeset
  1151
    else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))]
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1152
  end;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1153
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1154
val print_binds = Pretty.writeln o Pretty.chunks o pretty_binds;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1155
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1156
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1157
(* local theorems *)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1158
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1159
fun pretty_lthms ctxt =
20012
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1160
  let
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1161
    val props = FactIndex.props (fact_index_of ctxt);
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1162
    val facts =
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1163
      (if null props then I else cons ("unnamed", props))
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1164
      (NameSpace.extern_table (#1 (thms_of ctxt)));
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1165
  in
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1166
    if null facts andalso not (! verbose) then []
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1167
    else [Pretty.big_list "facts:" (map (pretty_fact ctxt) facts)]
b62836400a33 print_lthms: include unnamed facts from index;
wenzelm
parents: 20008
diff changeset
  1168
  end;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1169
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1170
val print_lthms = Pretty.writeln o Pretty.chunks o pretty_lthms;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1171
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1172
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1173
(* local contexts *)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1174
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1175
fun pretty_cases ctxt =
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1176
  let
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1177
    val prt_term = pretty_term ctxt;
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1178
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1179
    fun prt_let (xi, t) = Pretty.block
10818
37fa05a12459 tuned output;
wenzelm
parents: 10810
diff changeset
  1180
      [Pretty.quote (prt_term (Var (xi, Term.fastype_of t))), Pretty.str " =", Pretty.brk 1,
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1181
        Pretty.quote (prt_term t)];
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1182
13425
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
  1183
    fun prt_asm (a, ts) = Pretty.block (Pretty.breaks
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
  1184
      ((if a = "" then [] else [Pretty.str (a ^ ":")]) @ map (Pretty.quote o prt_term) ts));
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
  1185
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1186
    fun prt_sect _ _ _ [] = []
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1187
      | prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s ::
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19422
diff changeset
  1188
            flat (Library.separate sep (map (Library.single o prt) xs))))];
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1189
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1190
    fun prt_case (name, (fixes, (asms, (lets, cs)))) = Pretty.block (Pretty.fbreaks
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1191
      (Pretty.str (name ^ ":") ::
11915
df030220a2a8 print fixed names as plain strings;
wenzelm
parents: 11816
diff changeset
  1192
        prt_sect "fix" [] (Pretty.str o fst) fixes @
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1193
        prt_sect "let" [Pretty.str "and"] prt_let
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19422
diff changeset
  1194
          (map_filter (fn (xi, SOME t) => SOME (xi, t) | _ => NONE) lets) @
13425
119ae829ad9b support for split assumptions in cases (hyps vs. prems);
wenzelm
parents: 13415
diff changeset
  1195
        (if forall (null o #2) asms then []
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1196
          else prt_sect "assume" [Pretty.str "and"] prt_asm asms) @
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1197
        prt_sect "subcases:" [] (Pretty.str o fst) cs));
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1198
18476
49dde7b7b14a cases: main is_proper flag;
wenzelm
parents: 18428
diff changeset
  1199
    fun add_case (_, (_, false)) = I
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1200
      | add_case (name, (c as RuleCases.Case {fixes, ...}, true)) =
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1201
          cons (name, (fixes, #1 (case_result c ctxt)));
18476
49dde7b7b14a cases: main is_proper flag;
wenzelm
parents: 18428
diff changeset
  1202
    val cases = fold add_case (cases_of ctxt) [];
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1203
  in
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1204
    if null cases andalso not (! verbose) then []
18476
49dde7b7b14a cases: main is_proper flag;
wenzelm
parents: 18428
diff changeset
  1205
    else [Pretty.big_list "cases:" (map prt_case cases)]
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1206
  end;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1207
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1208
val print_cases = Pretty.writeln o Pretty.chunks o pretty_cases;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1209
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1210
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1211
(* core context *)
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1212
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
  1213
val prems_limit = ref ~1;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1214
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1215
fun pretty_ctxt ctxt =
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1216
  if ! prems_limit < 0 andalso not (! debug) then []
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1217
  else
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1218
    let
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1219
      val prt_term = pretty_term ctxt;
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1220
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1221
      (*structures*)
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1222
      val structs = LocalSyntax.structs_of (syntax_of ctxt);
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1223
      val prt_structs = if null structs then []
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1224
        else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 ::
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1225
          Pretty.commas (map Pretty.str structs))];
12093
1b890f1e0b4d syntax for structures;
wenzelm
parents: 12086
diff changeset
  1226
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1227
      (*fixes*)
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1228
      fun prt_fix (x, x') =
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1229
        if x = x' then Pretty.str x
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1230
        else Pretty.block [Pretty.str x, Pretty.str " =", Pretty.brk 1, prt_term (Syntax.free x')];
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1231
      val fixes =
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1232
        rev (filter_out ((can Name.dest_internal orf member (op =) structs) o #1)
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1233
          (Variable.fixes_of ctxt));
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1234
      val prt_fixes = if null fixes then []
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1235
        else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 ::
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1236
          Pretty.commas (map prt_fix fixes))];
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1237
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1238
      (*prems*)
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1239
      val prems = Assumption.prems_of ctxt;
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1240
      val len = length prems;
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
  1241
      val suppressed = len - ! prems_limit;
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1242
      val prt_prems = if null prems then []
20367
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
  1243
        else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @
ba1c262c7625 renamed map_theory to theory;
wenzelm
parents: 20330
diff changeset
  1244
          map (pretty_thm ctxt) (Library.drop (suppressed, prems)))];
20310
6cb47e95a74b normalized Proof.context/method type aliases;
wenzelm
parents: 20253
diff changeset
  1245
    in prt_structs @ prt_fixes @ prt_prems end;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1246
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1247
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1248
(* main context *)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1249
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1250
fun pretty_context ctxt =
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1251
  let
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1252
    val prt_term = pretty_term ctxt;
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1253
    val prt_typ = pretty_typ ctxt;
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1254
    val prt_sort = pretty_sort ctxt;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1255
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1256
    (*theory*)
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1257
    val pretty_thy = Pretty.block
17384
wenzelm
parents: 17360
diff changeset
  1258
      [Pretty.str "theory:", Pretty.brk 1, Context.pretty_thy (theory_of ctxt)];
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1259
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1260
    (*defaults*)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1261
    fun prt_atom prt prtT (x, X) = Pretty.block
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1262
      [prt x, Pretty.str " ::", Pretty.brk 1, prtT X];
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1263
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1264
    fun prt_var (x, ~1) = prt_term (Syntax.free x)
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1265
      | prt_var xi = prt_term (Syntax.var xi);
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1266
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1267
    fun prt_varT (x, ~1) = prt_typ (TFree (x, []))
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1268
      | prt_varT xi = prt_typ (TVar (xi, []));
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1269
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1270
    val prt_defT = prt_atom prt_var prt_typ;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1271
    val prt_defS = prt_atom prt_varT prt_sort;
16540
e3d61eff7c12 removed proof data (see Pure/context.ML);
wenzelm
parents: 16501
diff changeset
  1272
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
  1273
    val (types, sorts) = Variable.constraints_of ctxt;
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1274
  in
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1275
    verb single (K pretty_thy) @
18672
ac1a048ca7dd uniform handling of fixes: read/cert_vars, add_fixes(_i), body flag;
wenzelm
parents: 18619
diff changeset
  1276
    pretty_ctxt ctxt @
21728
906649272ba0 added read/pretty_term_abbrev, print_abbrevs;
wenzelm
parents: 21704
diff changeset
  1277
    verb (pretty_abbrevs false) (K ctxt) @
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1278
    verb pretty_binds (K ctxt) @
12057
9b1e67278f07 added pretty/print functions with context;
wenzelm
parents: 12048
diff changeset
  1279
    verb pretty_lthms (K ctxt) @
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1280
    verb pretty_cases (K ctxt) @
18609
b026652ede90 support nested cases;
wenzelm
parents: 18476
diff changeset
  1281
    verb single (fn () => Pretty.big_list "type constraints:" (map prt_defT (Vartab.dest types))) @
20163
08f2833ca433 Sign.infer_types: Name.context;
wenzelm
parents: 20101
diff changeset
  1282
    verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts)))
10810
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1283
  end;
619586bd854b apply_case: more robust handling of bounds;
wenzelm
parents: 10583
diff changeset
  1284
5819
5fff21d4ca3a Proof context information.
wenzelm
parents:
diff changeset
  1285
end;