src/Pure/Isar/local_defs.ML
author wenzelm
Tue, 03 Mar 2009 18:32:01 +0100
changeset 30223 24d975352879
parent 30211 556d1810cdad
child 30242 aea5d7fa7ef5
permissions -rw-r--r--
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify; minor tuning;
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
20306
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
     9
  val cert_def: Proof.context -> term -> (string * typ) * term
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    10
  val abs_def: term -> (string * typ) * term
20306
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    11
  val mk_def: Proof.context -> (string * term) list -> term list
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    12
  val expand: cterm list -> thm -> thm
20224
9c40a144ee0e moved basic assumption operations from structure ProofContext to Assumption;
wenzelm
parents: 20049
diff changeset
    13
  val def_export: Assumption.export
30211
556d1810cdad Thm.binding;
wenzelm
parents: 29581
diff changeset
    14
  val add_defs: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context ->
556d1810cdad Thm.binding;
wenzelm
parents: 29581
diff changeset
    15
    (term * (string * thm)) list * Proof.context
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29440
diff changeset
    16
  val add_def: (binding * mixfix) * term -> Proof.context -> (term * thm) * Proof.context
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29440
diff changeset
    17
  val fixed_abbrev: (binding * mixfix) * term -> Proof.context ->
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 26424
diff changeset
    18
    (term * term) * Proof.context
21591
5e0f2340caa7 added export;
wenzelm
parents: 21568
diff changeset
    19
  val export: Proof.context -> Proof.context -> thm -> thm list * thm
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
    20
  val export_cterm: Proof.context -> Proof.context -> cterm -> cterm * thm
21844
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
    21
  val trans_terms: Proof.context -> thm list -> thm
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
    22
  val trans_props: Proof.context -> thm list -> thm
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20980
diff changeset
    23
  val print_rules: Proof.context -> unit
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    24
  val defn_add: attribute
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    25
  val defn_del: attribute
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
    26
  val meta_rewrite_conv: Proof.context -> conv
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20980
diff changeset
    27
  val meta_rewrite_rule: Proof.context -> thm -> thm
20306
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    28
  val unfold: Proof.context -> thm list -> thm -> thm
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    29
  val unfold_goals: Proof.context -> thm list -> thm -> thm
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    30
  val unfold_tac: Proof.context -> thm list -> tactic
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    31
  val fold: Proof.context -> thm list -> thm -> thm
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    32
  val fold_tac: Proof.context -> thm list -> tactic
614b7e6c6276 normalized Proof.context/method type aliases;
wenzelm
parents: 20224
diff changeset
    33
  val derived_def: Proof.context -> bool -> term ->
20909
7132ab2b4621 simplified derived_def;
wenzelm
parents: 20887
diff changeset
    34
    ((string * typ) * term) * (Proof.context -> thm -> thm)
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    35
end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    36
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    37
structure LocalDefs: LOCAL_DEFS =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    38
struct
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    39
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    40
(** primitive definitions **)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
    41
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    42
(* prepare defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    43
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    44
fun cert_def ctxt eq =
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    45
  let
25025
17c845095a47 added add_def;
wenzelm
parents: 24985
diff changeset
    46
    fun err msg = cat_error msg ("The error(s) above occurred in definition: " ^
17c845095a47 added add_def;
wenzelm
parents: 24985
diff changeset
    47
      quote (Syntax.string_of_term ctxt eq));
18950
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
    48
    val ((lhs, _), eq') = eq
24981
4ec3f95190bf dest/cert_def: replaced Pretty.pp by explicit Proof.context;
wenzelm
parents: 24961
diff changeset
    49
      |> Sign.no_vars (Syntax.pp ctxt)
4ec3f95190bf dest/cert_def: replaced Pretty.pp by explicit Proof.context;
wenzelm
parents: 24961
diff changeset
    50
      |> PrimitiveDefs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true)
18950
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
    51
      handle TERM (msg, _) => err msg | ERROR msg => err msg;
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
    52
  in (Term.dest_Free (Term.head_of lhs), eq') end;
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    53
24261
dd31811bdf46 PrimitiveDefs.dest/abs_def;
wenzelm
parents: 24039
diff changeset
    54
val abs_def = PrimitiveDefs.abs_def #>> Term.dest_Free;
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    55
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    56
fun mk_def ctxt args =
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    57
  let
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    58
    val (xs, rhss) = split_list args;
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    59
    val (bind, _) = ProofContext.bind_fixes xs ctxt;
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    60
    val lhss = map (fn (x, rhs) => bind (Free (x, Term.fastype_of rhs))) args;
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    61
  in map Logic.mk_equals (lhss ~~ rhss) end;
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    62
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    63
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    64
(* export defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    65
20021
815393c02db9 def_export: Drule.generalize;
wenzelm
parents: 19897
diff changeset
    66
val head_of_def =
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    67
  #1 o 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
    68
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    69
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    70
(*
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    71
  [x, x == a]
18896
efd9d44a0bdb tuned comments;
wenzelm
parents: 18875
diff changeset
    72
       :
efd9d44a0bdb tuned comments;
wenzelm
parents: 18875
diff changeset
    73
      B x
efd9d44a0bdb tuned comments;
wenzelm
parents: 18875
diff changeset
    74
  -----------
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    75
      B a
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
    76
*)
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    77
fun expand defs =
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    78
  Drule.implies_intr_list defs
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    79
  #> Drule.generalize ([], map (head_of_def o Thm.term_of) defs)
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    80
  #> funpow (length defs) (fn th => Drule.reflexive_thm RS th);
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    81
21801
c77bc21b3eef tuned expand_term;
wenzelm
parents: 21708
diff changeset
    82
val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of);
21684
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    83
e8c135b166b3 added expand;
wenzelm
parents: 21591
diff changeset
    84
fun def_export _ defs = (expand defs, expand_term defs);
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    85
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    86
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    87
(* add defs *)
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    88
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    89
fun add_defs defs ctxt =
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    90
  let
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 26424
diff changeset
    91
    val ((bvars, mxs), specs) = defs |> split_list |>> split_list;
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 26424
diff changeset
    92
    val ((bfacts, atts), rhss) = specs |> split_list |>> split_list;
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30211
diff changeset
    93
    val xs = map Binding.name_of bvars;
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30211
diff changeset
    94
    val names = map2 (Binding.map_name o Thm.def_name_optional) xs bfacts;
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    95
    val eqs = mk_def ctxt (xs ~~ rhss);
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
    96
    val lhss = map (fst o Logic.dest_equals) eqs;
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    97
  in
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
    98
    ctxt
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 26424
diff changeset
    99
    |> ProofContext.add_fixes_i (map2 (fn x => fn mx => (x, NONE, mx)) bvars mxs) |> #2
20980
e4fd72aecd03 add_defs: declare terms;
wenzelm
parents: 20909
diff changeset
   100
    |> fold Variable.declare_term eqs
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
   101
    |> ProofContext.add_assms_i def_export
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 26424
diff changeset
   102
      (map2 (fn a => fn eq => (a, [(eq, [])])) (names ~~ atts) eqs)
20887
ec9a1bb086da replaced add_def by more elaborate add_defs;
wenzelm
parents: 20306
diff changeset
   103
    |>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   104
  end;
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   105
25104
26b9a7db3386 tuned interfaces;
wenzelm
parents: 25025
diff changeset
   106
fun add_def (var, rhs) ctxt =
30211
556d1810cdad Thm.binding;
wenzelm
parents: 29581
diff changeset
   107
  let val ([(lhs, (_, th))], ctxt') = add_defs [(var, (Thm.empty_binding, rhs))] ctxt
25104
26b9a7db3386 tuned interfaces;
wenzelm
parents: 25025
diff changeset
   108
  in ((lhs, th), ctxt') end;
25025
17c845095a47 added add_def;
wenzelm
parents: 24985
diff changeset
   109
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   110
25119
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   111
(* fixed_abbrev *)
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   112
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   113
fun fixed_abbrev ((x, mx), rhs) ctxt =
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   114
  let
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   115
    val T = Term.fastype_of rhs;
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   116
    val ([x'], ctxt') = ctxt
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   117
      |> Variable.declare_term rhs
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   118
      |> ProofContext.add_fixes_i [(x, SOME T, mx)];
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   119
    val lhs = Free (x', T);
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   120
    val _ = cert_def ctxt' (Logic.mk_equals (lhs, rhs));
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   121
    fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]);
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   122
    val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt';
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   123
  in ((lhs, rhs), ctxt'') end;
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   124
192105867f25 added fixed_abbrev;
wenzelm
parents: 25104
diff changeset
   125
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   126
(* specific export -- result based on educated guessing *)
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   127
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   128
(*
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   129
  [xs, xs == as]
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   130
        :
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   131
       B xs
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   132
  --------------
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   133
       B as
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   134
*)
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   135
fun export inner outer =    (*beware of closure sizes*)
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   136
  let
22671
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   137
    val exp = Assumption.export false inner outer;
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   138
    val prems = Assumption.prems_of inner;
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   139
  in fn th =>
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   140
    let
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   141
      val th' = exp th;
22691
290454649b8c Thm.fold_terms;
wenzelm
parents: 22671
diff changeset
   142
      val still_fixed = map #1 (Thm.fold_terms Term.add_frees th' []);
22671
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   143
      val defs = prems |> filter_out (fn prem =>
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   144
        (case try (head_of_def o Thm.prop_of) prem of
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   145
          SOME x => member (op =) still_fixed x
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   146
        | NONE => true));
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   147
    in (map Drule.abs_def defs, th') end
3c62305fbee6 tuned signature;
wenzelm
parents: 22568
diff changeset
   148
  end;
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   149
24985
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   150
(*
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   151
  [xs, xs == as]
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   152
        :
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   153
     TERM b xs
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   154
  --------------  and  --------------
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   155
     TERM b as          b xs == b as
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   156
*)
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   157
fun export_cterm inner outer ct =
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   158
  let val (defs, ct') = export inner outer (Drule.mk_term ct) ||> Drule.dest_term
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   159
  in (ct', MetaSimplifier.rewrite true defs ct) end;
0e5177e2c076 added export_cterm;
wenzelm
parents: 24981
diff changeset
   160
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   161
21844
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   162
(* basic transitivity reasoning -- modulo beta-eta *)
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   163
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   164
local
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   165
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   166
val is_trivial = Pattern.aeconv o Logic.dest_equals o Thm.prop_of;
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   167
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   168
fun trans_list _ _ [] = raise Empty
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   169
  | trans_list trans ctxt (th :: raw_eqs) =
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   170
      (case filter_out is_trivial raw_eqs of
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   171
        [] => th
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   172
      | eqs =>
22568
ed7aa5a350ef renamed Variable.import to import_thms (avoid clash with Alice keywords);
wenzelm
parents: 21844
diff changeset
   173
          let val ((_, th' :: eqs'), ctxt') = Variable.import_thms true (th :: eqs) ctxt
21844
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   174
          in singleton (Variable.export ctxt' ctxt) (fold trans eqs' th') end);
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   175
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   176
in
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   177
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   178
val trans_terms = trans_list (fn eq2 => fn eq1 => eq2 COMP (eq1 COMP Drule.transitive_thm));
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   179
val trans_props = trans_list (fn eq => fn th => th COMP (eq COMP Drule.equal_elim_rule1));
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   180
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   181
end;
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   182
e10b8bd7a886 added trans_terms/props;
wenzelm
parents: 21801
diff changeset
   183
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   184
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   185
(** defived definitions **)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   186
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   187
(* transformation rules *)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   188
18859
75248f8badc9 export meta_rewrite_rule;
wenzelm
parents: 18840
diff changeset
   189
structure Rules = GenericDataFun
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   190
(
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   191
  type T = thm list;
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   192
  val empty = []
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   193
  val extend = I;
24039
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23541
diff changeset
   194
  fun merge _ = Thm.merge_thms;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   195
);
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   196
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22691
diff changeset
   197
fun print_rules ctxt =
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22691
diff changeset
   198
  Pretty.writeln (Pretty.big_list "definitional transformations:"
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22691
diff changeset
   199
    (map (ProofContext.pretty_thm ctxt) (Rules.get (Context.Proof ctxt))));
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   200
24039
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23541
diff changeset
   201
val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm);
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23541
diff changeset
   202
val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm);
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   203
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   204
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   205
(* meta rewrite rules *)
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   206
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   207
fun meta_rewrite_conv ctxt =
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   208
  MetaSimplifier.rewrite_cterm (false, false, false) (K (K NONE))
26424
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 25119
diff changeset
   209
    (MetaSimplifier.context ctxt MetaSimplifier.empty_ss
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 25119
diff changeset
   210
      addeqcongs [Drule.equals_cong]    (*protect meta-level equality*)
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 25119
diff changeset
   211
      addsimps (Rules.get (Context.Proof ctxt)));
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   212
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   213
val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   214
18875
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   215
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   216
(* rewriting with object-level rules *)
853fa34047a4 (un)fold: no raw flag;
wenzelm
parents: 18859
diff changeset
   217
21506
b2a673894ce5 prefer Proof.context over Context.generic;
wenzelm
parents: 20980
diff changeset
   218
fun meta f ctxt = f o map (meta_rewrite_rule ctxt);
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   219
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21699
diff changeset
   220
val unfold       = meta MetaSimplifier.rewrite_rule;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21699
diff changeset
   221
val unfold_goals = meta MetaSimplifier.rewrite_goals_rule;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21699
diff changeset
   222
val unfold_tac   = meta MetaSimplifier.rewrite_goals_tac;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21699
diff changeset
   223
val fold         = meta MetaSimplifier.fold_rule;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21699
diff changeset
   224
val fold_tac     = meta MetaSimplifier.fold_goals_tac;
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   225
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   226
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   227
(* derived defs -- potentially within the object-logic *)
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   228
18950
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
   229
fun derived_def ctxt conditional prop =
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   230
  let
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   231
    val ((c, T), rhs) = prop
20049
f48c4a3a34bc Goal.prove: context;
wenzelm
parents: 20021
diff changeset
   232
      |> Thm.cterm_of (ProofContext.theory_of ctxt)
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   233
      |> meta_rewrite_conv ctxt
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   234
      |> (snd o Logic.dest_equals o Thm.prop_of)
21568
a8070c4b6d43 simplified '?' operator;
wenzelm
parents: 21506
diff changeset
   235
      |> conditional ? Logic.strip_imp_concl
18950
053e830c25ad cert_def: use Logic.dest_def;
wenzelm
parents: 18896
diff changeset
   236
      |> (abs_def o #2 o cert_def ctxt);
20909
7132ab2b4621 simplified derived_def;
wenzelm
parents: 20887
diff changeset
   237
    fun prove ctxt' def =
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   238
      let
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   239
        val frees = Term.fold_aterms (fn Free (x, _) =>
20909
7132ab2b4621 simplified derived_def;
wenzelm
parents: 20887
diff changeset
   240
          if Variable.is_fixed ctxt' x then I else insert (op =) x | _ => I) prop [];
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   241
      in
20909
7132ab2b4621 simplified derived_def;
wenzelm
parents: 20887
diff changeset
   242
        Goal.prove ctxt' frees [] prop (K (ALLGOALS
23541
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   243
          (CONVERSION (meta_rewrite_conv ctxt') THEN'
f8c5e218e4d8 exported meta_rewrite_conv;
wenzelm
parents: 22900
diff changeset
   244
            MetaSimplifier.rewrite_goal_tac [def] THEN'
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21699
diff changeset
   245
            resolve_tac [Drule.reflexive_thm])))
18840
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   246
        handle ERROR msg => cat_error msg "Failed to prove definitional specification."
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   247
      end;
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   248
  in (((c, T), rhs), prove) end;
ce16e2bad548 added attributes defn_add/del;
wenzelm
parents: 18830
diff changeset
   249
18830
34b51dcdc570 Basic operations on local definitions.
wenzelm
parents:
diff changeset
   250
end;