src/Pure/Isar/local_defs.ML
author wenzelm
Thu, 26 Aug 2021 14:45:19 +0200
changeset 74200 17090e27aae9
parent 70314 6d6839a948cf
child 74266 612b7e0d6721
permissions -rw-r--r--
more scalable data structure (but: rarely used with > 5 arguments);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/local_defs.ML
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     3
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
     4
Local definitions.
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     5
*)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     6
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     7
signature LOCAL_DEFS =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
     8
sig
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
     9
  val cert_def: Proof.context -> (string -> Position.T list) -> term -> (string * typ) * term
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    10
  val abs_def: term -> (string * typ) * term
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    11
  val expand: cterm list -> thm -> thm
20224
9c40a144ee0e moved basic assumption operations from structure ProofContext to Assumption;
wenzelm
parents: 20049
diff changeset
    12
  val def_export: Assumption.export
63344
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
    13
  val define: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context ->
30211
556d1810cdad Thm.binding;
wenzelm
parents: 29581
diff changeset
    14
    (term * (string * thm)) list * Proof.context
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29440
diff changeset
    15
  val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 26424
diff changeset
    16
    (term * term) * Proof.context
45407
a83574606719 more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
wenzelm
parents: 45406
diff changeset
    17
  val export: Proof.context -> Proof.context -> thm -> (thm list * thm list) * thm
a83574606719 more specific treatment of defines/assumes -- avoid normalizing defs by themselves (NB: locale specifications and Local_Theory.define may lead to arbitrary mixture);
wenzelm
parents: 45406
diff changeset
    18
  val export_cterm: Proof.context -> Proof.context -> cterm -> (thm list * thm list) * cterm
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
    19
  val contract: Proof.context -> thm list -> cterm -> thm -> thm
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20980
diff changeset
    20
  val print_rules: Proof.context -> unit
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    21
  val defn_add: attribute
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    22
  val defn_del: attribute
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
    23
  val meta_rewrite_conv: Proof.context -> conv
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20980
diff changeset
    24
  val meta_rewrite_rule: Proof.context -> thm -> thm
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
    25
  val abs_def_rule: Proof.context -> thm -> thm
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
    26
  val unfold_abs_def: bool Config.T
20306
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    27
  val unfold: Proof.context -> thm list -> thm -> thm
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    28
  val unfold_goals: Proof.context -> thm list -> thm -> thm
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    29
  val unfold_tac: Proof.context -> thm list -> tactic
63169
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
    30
  val unfold0: Proof.context -> thm list -> thm -> thm
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
    31
  val unfold0_goals: Proof.context -> thm list -> thm -> thm
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
    32
  val unfold0_tac: Proof.context -> thm list -> tactic
20306
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    33
  val fold: Proof.context -> thm list -> thm -> thm
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    34
  val fold_tac: Proof.context -> thm list -> tactic
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    35
  val derived_def: Proof.context -> (string -> Position.T list) -> {conditional: bool} ->
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    36
    term -> ((string * typ) * term) * (Proof.context -> thm -> thm)
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    37
end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    38
35624
c4e29a0bb8c1 modernized structure Local_Defs;
wenzelm
parents: 33519
diff changeset
    39
structure Local_Defs: LOCAL_DEFS =
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    40
struct
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    41
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    42
(** primitive definitions **)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    43
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    44
(* prepare defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    45
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    46
fun cert_def ctxt get_pos eq =
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    47
  let
40242
bb433b0668b8 tuned messages;
wenzelm
parents: 39133
diff changeset
    48
    fun err msg =
bb433b0668b8 tuned messages;
wenzelm
parents: 39133
diff changeset
    49
      cat_error msg ("The error(s) above occurred in definition:\n" ^
bb433b0668b8 tuned messages;
wenzelm
parents: 39133
diff changeset
    50
        quote (Syntax.string_of_term ctxt eq));
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    51
    val ((lhs, _), args, eq') = eq
39133
70d3915c92f0 pretty printing: prefer regular Proof.context over Pretty.pp, which is mostly for special bootstrap purposes involving theory merge, for example;
wenzelm
parents: 35739
diff changeset
    52
      |> Sign.no_vars ctxt
63042
741263be960e more rigid check of lhs;
wenzelm
parents: 63038
diff changeset
    53
      |> Primitive_Defs.dest_def ctxt
741263be960e more rigid check of lhs;
wenzelm
parents: 63038
diff changeset
    54
        {check_head = Term.is_Free,
741263be960e more rigid check of lhs;
wenzelm
parents: 63038
diff changeset
    55
         check_free_lhs = not o Variable.is_fixed ctxt,
741263be960e more rigid check of lhs;
wenzelm
parents: 63038
diff changeset
    56
         check_free_rhs = if Variable.is_body ctxt then K true else Variable.is_fixed ctxt,
741263be960e more rigid check of lhs;
wenzelm
parents: 63038
diff changeset
    57
         check_tfree = K true}
18950
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
    58
      handle TERM (msg, _) => err msg | ERROR msg => err msg;
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    59
    val _ =
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    60
      Context_Position.reports ctxt
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
    61
        (maps (fn Free (x, _) => Syntax_Phases.reports_of_scope (get_pos x) | _ => []) args);
18950
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
    62
  in (Term.dest_Free (Term.head_of lhs), eq') end;
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    63
33385
fb2358edcfb6 modernized structure Primitive_Defs;
wenzelm
parents: 32091
diff changeset
    64
val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free;
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    65
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    66
fun mk_def ctxt args =
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    67
  let
42501
2b8c34f53388 eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents: 42496
diff changeset
    68
    val (bs, rhss) = split_list args;
2b8c34f53388 eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents: 42496
diff changeset
    69
    val Ts = map Term.fastype_of rhss;
63344
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
    70
    val (xs, _) = ctxt
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
    71
      |> Context_Position.set_visible false
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
    72
      |> Proof_Context.add_fixes (map2 (fn b => fn T => (b, SOME T, NoSyn)) bs Ts);
42501
2b8c34f53388 eliminated slightly odd Proof_Context.bind_fixes;
wenzelm
parents: 42496
diff changeset
    73
    val lhss = ListPair.map Free (xs, Ts);
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    74
  in map Logic.mk_equals (lhss ~~ rhss) end;
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    75
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    76
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    77
(* export defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    78
20021
815393c02db9 def_export: Drule.generalize;
wenzelm
parents: 19897
diff changeset
    79
val head_of_def =
45406
b24ecaa46edb clarified Local_Defs.export: avoid costly still_fixed test, return all defs;
wenzelm
parents: 42501
diff changeset
    80
  Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body;
20021
815393c02db9 def_export: Drule.generalize;
wenzelm
parents: 19897
diff changeset
    81
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    82
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    83
(*
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67627
diff changeset
    84
  [x, x \<equiv> a]
18896
efd9d44a0bdb tuned comments;
wenzelm
parents: 18875
diff changeset
    85
       :
efd9d44a0bdb tuned comments;
wenzelm
parents: 18875
diff changeset
    86
      B x
efd9d44a0bdb tuned comments;
wenzelm
parents: 18875
diff changeset
    87
  -----------
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    88
      B a
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    89
*)
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    90
fun expand defs =
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    91
  Drule.implies_intr_list defs
74200
17090e27aae9 more scalable data structure (but: rarely used with > 5 arguments);
wenzelm
parents: 70314
diff changeset
    92
  #> Drule.generalize
17090e27aae9 more scalable data structure (but: rarely used with > 5 arguments);
wenzelm
parents: 70314
diff changeset
    93
      (Symtab.empty, fold (Symtab.insert_set o #1 o head_of_def o Thm.term_of) defs Symtab.empty)
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    94
  #> funpow (length defs) (fn th => Drule.reflexive_thm RS th);
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    95
21801
c77bc21b3eef tuned expand_term;
wenzelm
parents: 21708
diff changeset
    96
val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of);
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    97
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    98
fun def_export _ defs = (expand defs, expand_term defs);
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    99
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   100
63344
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
   101
(* define *)
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   102
63344
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
   103
fun define defs ctxt =
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   104
  let
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 42494
diff changeset
   105
    val ((xs, mxs), specs) = defs |> split_list |>> split_list;
49748
a346daa8a1f4 clarified Local_Defs.add_def(s): refrain from hard-wiring Thm.def_binding_optional;
wenzelm
parents: 47296
diff changeset
   106
    val (bs, rhss) = specs |> split_list;
63344
c9910404cc8a tuned signature;
wenzelm
parents: 63221
diff changeset
   107
    val eqs = mk_def ctxt (xs ~~ rhss);
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
   108
    val lhss = map (fst o Logic.dest_equals) eqs;
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   109
  in
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   110
    ctxt
42496
65ec88b369fd more precise positions via binding;
wenzelm
parents: 42494
diff changeset
   111
    |> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2
20980
e4fd72aecd03 add_defs: declare terms;
wenzelm
parents: 20909
diff changeset
   112
    |> fold Variable.declare_term eqs
60377
234b7da8542e tuned signature;
wenzelm
parents: 60240
diff changeset
   113
    |> Proof_Context.add_assms def_export (map2 (fn b => fn eq => (b, [(eq, [])])) bs eqs)
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
   114
    |>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   115
  end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   116
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   117
25119
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   118
(* fixed_abbrev *)
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   119
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   120
fun fixed_abbrev ((x, mx), rhs) ctxt =
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   121
  let
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   122
    val T = Term.fastype_of rhs;
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   123
    val ([x'], ctxt') = ctxt
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   124
      |> Variable.declare_term rhs
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42357
diff changeset
   125
      |> Proof_Context.add_fixes [(x, SOME T, mx)];
25119
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   126
    val lhs = Free (x', T);
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
   127
    val _ = cert_def ctxt' (K []) (Logic.mk_equals (lhs, rhs));
25119
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   128
    fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]);
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   129
    val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt';
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   130
  in ((lhs, rhs), ctxt'') end;
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   131
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   132
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   133
(* specific export -- result based on educated guessing *)
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   134
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   135
(*
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67627
diff changeset
   136
  [xs, xs \<equiv> as]
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   137
        :
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   138
       B xs
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   139
  --------------
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   140
       B as
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   141
*)
60823
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   142
fun export inner outer th =
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   143
  let
60823
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   144
    val defs_asms =
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   145
      Assumption.local_assms_of inner outer
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   146
      |> filter_out (Drule.is_sort_constraint o Thm.term_of)
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   147
      |> map (Thm.assume #> (fn asm =>
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   148
        (case try (head_of_def o Thm.prop_of) asm of
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   149
          NONE => (asm, false)
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   150
        | SOME x =>
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   151
            let val t = Free x in
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   152
              (case try (Assumption.export_term inner outer) t of
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   153
                NONE => (asm, false)
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   154
              | SOME u =>
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   155
                  if t aconv u then (asm, false)
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   156
                  else (Drule.abs_def (Variable.gen_all outer asm), true))
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   157
            end)));
b41478500473 clarified Variable.gen_all;
wenzelm
parents: 60822
diff changeset
   158
  in (apply2 (map #1) (List.partition #2 defs_asms), Assumption.export false inner outer th) end;
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   159
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   160
(*
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67627
diff changeset
   161
  [xs, xs \<equiv> as]
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   162
        :
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   163
     TERM b xs
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   164
  --------------  and  --------------
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67627
diff changeset
   165
     TERM b as          b xs \<equiv> b as
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   166
*)
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   167
fun export_cterm inner outer ct =
35717
1856c0172cf2 more basic Local_Defs.export_cterm;
wenzelm
parents: 35624
diff changeset
   168
  export inner outer (Drule.mk_term ct) ||> Drule.dest_term;
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   169
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   170
fun contract ctxt defs ct th =
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   171
  th COMP (Raw_Simplifier.rewrite ctxt true defs ct COMP_INCR Drule.equal_elim_rule2);
21844
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   172
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   173
47237
b61a11dea74c more precise Local_Defs.expand wrt. *local* prems only;
wenzelm
parents: 47060
diff changeset
   174
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   175
(** defived definitions **)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   176
51585
fcd5af4aac2b added 'print_defn_rules' command;
wenzelm
parents: 51584
diff changeset
   177
(* transformation via rewrite rules *)
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   178
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33385
diff changeset
   179
structure Rules = Generic_Data
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   180
(
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   181
  type T = thm list;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33385
diff changeset
   182
  val empty = [];
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   183
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33385
diff changeset
   184
  val merge = Thm.merge_thms;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   185
);
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   186
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22691
diff changeset
   187
fun print_rules ctxt =
51585
fcd5af4aac2b added 'print_defn_rules' command;
wenzelm
parents: 51584
diff changeset
   188
  Pretty.writeln (Pretty.big_list "definitional rewrite rules:"
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61091
diff changeset
   189
    (map (Thm.pretty_thm_item ctxt) (Rules.get (Context.Proof ctxt))));
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   190
61091
2b7ef52a4ea9 trim context for persistent storage;
wenzelm
parents: 60823
diff changeset
   191
val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm o Thm.trim_context);
67627
5cca859b2d2e trim context of persistent data;
wenzelm
parents: 64556
diff changeset
   192
val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm);
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   193
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   194
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   195
(* meta rewrite rules *)
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   196
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   197
fun meta_rewrite_conv ctxt =
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 40242
diff changeset
   198
  Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE))
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 63169
diff changeset
   199
    (ctxt
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 63169
diff changeset
   200
      |> Raw_Simplifier.init_simpset (Rules.get (Context.Proof ctxt))
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45407
diff changeset
   201
      |> Raw_Simplifier.add_eqcong Drule.equals_cong);    (*protect meta-level equality*)
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   202
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   203
val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   204
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   205
fun abs_def_rule ctxt = meta_rewrite_rule ctxt #> Drule.abs_def;
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   206
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   207
63169
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   208
(* unfold object-level rules *)
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   209
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 67721
diff changeset
   210
val unfold_abs_def = Config.declare_bool ("unfold_abs_def", \<^here>) (K true);
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   211
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   212
local
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   213
63169
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   214
fun gen_unfold rewrite ctxt rews =
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   215
  let val meta_rews = map (meta_rewrite_rule ctxt) rews in
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   216
    if Config.get ctxt unfold_abs_def then
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   217
      rewrite ctxt meta_rews #>
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   218
      rewrite ctxt (map (perhaps (try Drule.abs_def)) meta_rews)
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   219
    else rewrite ctxt meta_rews
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   220
  end;
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   221
63169
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   222
val no_unfold_abs_def = Config.put unfold_abs_def false;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   223
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   224
in
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   225
63169
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   226
val unfold = gen_unfold Raw_Simplifier.rewrite_rule;
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   227
val unfold_goals = gen_unfold Raw_Simplifier.rewrite_goals_rule;
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   228
val unfold_tac = PRIMITIVE oo unfold_goals;
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   229
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   230
val unfold0 = unfold o no_unfold_abs_def;
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   231
val unfold0_goals = unfold_goals o no_unfold_abs_def;
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   232
val unfold0_tac = unfold_tac o no_unfold_abs_def;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   233
63169
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   234
end
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   235
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   236
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   237
(* fold object-level rules *)
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   238
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   239
fun fold ctxt rews = Raw_Simplifier.fold_rule ctxt (map (meta_rewrite_rule ctxt) rews);
d36c7dc40000 clarified "unfold" operations;
wenzelm
parents: 63068
diff changeset
   240
fun fold_tac ctxt rews = Raw_Simplifier.fold_goals_tac ctxt (map (meta_rewrite_rule ctxt) rews);
63068
8b9401bfd9fd unfold is subject to unfold_abs_def (still inactive);
wenzelm
parents: 63042
diff changeset
   241
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   242
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   243
(* derived defs -- potentially within the object-logic *)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   244
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
   245
fun derived_def ctxt get_pos {conditional} prop =
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   246
  let
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   247
    val ((c, T), rhs) = prop
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59573
diff changeset
   248
      |> Thm.cterm_of ctxt
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   249
      |> meta_rewrite_conv ctxt
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   250
      |> (snd o Logic.dest_equals o Thm.prop_of)
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   251
      |> conditional ? Logic.strip_imp_concl
63395
734723445a8c PIDE reports of implicit variable scope;
wenzelm
parents: 63344
diff changeset
   252
      |> (abs_def o #2 o cert_def ctxt get_pos);
70306
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   253
    fun prove def_ctxt0 def =
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   254
      let
70308
7f568724d67e clarified signature;
wenzelm
parents: 70306
diff changeset
   255
        val def_ctxt = Proof_Context.augment prop def_ctxt0;
70306
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   256
        val def_thm =
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   257
          Goal.prove def_ctxt [] [] prop
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   258
            (fn {context = goal_ctxt, ...} =>
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   259
              ALLGOALS
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   260
                (CONVERSION (meta_rewrite_conv goal_ctxt) THEN'
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   261
                  rewrite_goal_tac goal_ctxt [def] THEN'
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   262
                  resolve_tac goal_ctxt [Drule.reflexive_thm]))
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   263
          handle ERROR msg => cat_error msg "Failed to prove definitional specification";
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   264
      in
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   265
        def_thm
70314
6d6839a948cf proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
wenzelm
parents: 70308
diff changeset
   266
        |> singleton (Proof_Context.export def_ctxt def_ctxt0)
70306
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   267
        |> Drule.zero_var_indexes
61621dc439d4 clarified context: prefer abstract Variable.auto_fixes;
wenzelm
parents: 70305
diff changeset
   268
      end;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   269
  in (((c, T), rhs), prove) end;
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   270
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   271
end;