src/Pure/raw_simplifier.ML
author wenzelm
Thu, 12 Dec 2013 22:38:25 +0100
changeset 54731 384ac33802b0
parent 54729 c5cd7a58cf2d
child 54742 7a86358a3c0b
permissions -rw-r--r--
clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
     1
(*  Title:      Pure/raw_simplifier.ML
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 28839
diff changeset
     2
    Author:     Tobias Nipkow and Stefan Berghofer, TU Muenchen
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
     3
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
     4
Higher-order Simplification.
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
     5
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
     6
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
     7
infix 4
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45405
diff changeset
     8
  addsimps delsimps addsimprocs delsimprocs
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
     9
  setloop addloop delloop
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
    10
  setSSolver addSSolver setSolver addSolver;
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
    11
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
    12
signature BASIC_RAW_SIMPLIFIER =
11672
8e75b78f33f3 full_rewrite_cterm_aux (see also tactic.ML);
wenzelm
parents: 11505
diff changeset
    13
sig
41227
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
    14
  val simp_depth_limit: int Config.T
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
    15
  val simp_trace_depth_limit: int Config.T
40878
7695e4de4d86 renamed trace_simp to simp_trace, and debug_simp to simp_debug;
wenzelm
parents: 39163
diff changeset
    16
  val simp_debug: bool Config.T
7695e4de4d86 renamed trace_simp to simp_trace, and debug_simp to simp_debug;
wenzelm
parents: 39163
diff changeset
    17
  val simp_trace: bool Config.T
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
    18
  type cong_name = bool * string
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    19
  type rrule
16807
730cace0ae48 export eq_rrule;
wenzelm
parents: 16665
diff changeset
    20
  val eq_rrule: rrule * rrule -> bool
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    21
  type proc
17614
37ee526db497 added mk_solver';
wenzelm
parents: 17496
diff changeset
    22
  type solver
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    23
  val mk_solver: string -> (Proof.context -> int -> tactic) -> solver
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    24
  type simpset
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    25
  val empty_ss: simpset
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    26
  val merge_ss: simpset * simpset -> simpset
30356
36d0e00af606 added dest_ss;
wenzelm
parents: 30336
diff changeset
    27
  val dest_ss: simpset ->
36d0e00af606 added dest_ss;
wenzelm
parents: 30336
diff changeset
    28
   {simps: (string * thm) list,
36d0e00af606 added dest_ss;
wenzelm
parents: 30336
diff changeset
    29
    procs: (string * cterm list) list,
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
    30
    congs: (cong_name * thm) list,
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
    31
    weak_congs: cong_name list,
30356
36d0e00af606 added dest_ss;
wenzelm
parents: 30336
diff changeset
    32
    loopers: string list,
36d0e00af606 added dest_ss;
wenzelm
parents: 30336
diff changeset
    33
    unsafe_solvers: string list,
36d0e00af606 added dest_ss;
wenzelm
parents: 30336
diff changeset
    34
    safe_solvers: string list}
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    35
  type simproc
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
    36
  val eq_simproc: simproc * simproc -> bool
45290
f599ac41e7f5 tuned signature -- refined terminology;
wenzelm
parents: 44058
diff changeset
    37
  val transform_simproc: morphism -> simproc -> simproc
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
    38
  val make_simproc: {name: string, lhss: cterm list,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    39
    proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> simproc
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    40
  val mk_simproc: string -> cterm list -> (Proof.context -> term -> thm option) -> simproc
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    41
  val simpset_of: Proof.context -> simpset
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    42
  val put_simpset: simpset -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    43
  val global_context: theory -> simpset -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    44
  val simpset_map: Proof.context -> (Proof.context -> Proof.context) -> simpset -> simpset
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    45
  val map_theory_simpset: (Proof.context -> Proof.context) -> theory -> theory
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    46
  val empty_simpset: Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    47
  val clear_simpset: Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    48
  val addsimps: Proof.context * thm list -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    49
  val delsimps: Proof.context * thm list -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    50
  val addsimprocs: Proof.context * simproc list -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    51
  val delsimprocs: Proof.context * simproc list -> Proof.context
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
    52
  val setloop: Proof.context * (Proof.context -> int -> tactic) -> Proof.context
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
    53
  val addloop: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    54
  val delloop: Proof.context * string -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    55
  val setSSolver: Proof.context * solver -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    56
  val addSSolver: Proof.context * solver -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    57
  val setSolver: Proof.context * solver -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    58
  val addSolver: Proof.context * solver -> Proof.context
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    59
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    60
  val rewrite_rule: thm list -> thm -> thm
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    61
  val rewrite_goals_rule: thm list -> thm -> thm
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    62
  val rewrite_goals_tac: thm list -> tactic
23536
60a1672e298e moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
wenzelm
parents: 23221
diff changeset
    63
  val rewrite_goal_tac: thm list -> int -> tactic
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    64
  val prune_params_tac: tactic
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    65
  val fold_rule: thm list -> thm -> thm
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    66
  val fold_goals_tac: thm list -> tactic
30552
58db56278478 provide Simplifier.norm_hhf(_protect) as regular simplifier operation;
wenzelm
parents: 30356
diff changeset
    67
  val norm_hhf: thm -> thm
58db56278478 provide Simplifier.norm_hhf(_protect) as regular simplifier operation;
wenzelm
parents: 30356
diff changeset
    68
  val norm_hhf_protect: thm -> thm
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
    69
end;
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
    70
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
    71
signature RAW_SIMPLIFIER =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
    72
sig
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
    73
  include BASIC_RAW_SIMPLIFIER
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
    74
  exception SIMPLIFIER of string * thm
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
    75
  type trace_ops
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
    76
  val set_trace_ops: trace_ops -> theory -> theory
30336
efd1bec4630a renamed rep_ss to MetaSimplifier.internal_ss;
wenzelm
parents: 30318
diff changeset
    77
  val internal_ss: simpset ->
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
    78
   {congs: (cong_name * thm) list * cong_name list,
30336
efd1bec4630a renamed rep_ss to MetaSimplifier.internal_ss;
wenzelm
parents: 30318
diff changeset
    79
    procs: proc Net.net,
efd1bec4630a renamed rep_ss to MetaSimplifier.internal_ss;
wenzelm
parents: 30318
diff changeset
    80
    mk_rews:
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    81
     {mk: Proof.context -> thm -> thm list,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    82
      mk_cong: Proof.context -> thm -> thm,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    83
      mk_sym: Proof.context -> thm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    84
      mk_eq_True: Proof.context -> thm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    85
      reorient: Proof.context -> term list -> term -> term -> bool},
30336
efd1bec4630a renamed rep_ss to MetaSimplifier.internal_ss;
wenzelm
parents: 30318
diff changeset
    86
    termless: term * term -> bool,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    87
    subgoal_tac: Proof.context -> int -> tactic,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    88
    loop_tacs: (string * (Proof.context -> int -> tactic)) list,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
    89
    solvers: solver list * solver list}
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    90
  val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    91
  val prems_of: Proof.context -> thm list
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    92
  val add_simp: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    93
  val del_simp: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    94
  val add_eqcong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    95
  val del_eqcong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    96
  val add_cong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    97
  val del_cong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    98
  val mksimps: Proof.context -> thm -> thm list
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    99
  val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   100
  val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   101
  val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   102
  val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   103
  val set_termless: (term * term -> bool) -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   104
  val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   105
  val solver: Proof.context -> solver -> int -> tactic
39163
4d701c0388c3 more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents: 39116
diff changeset
   106
  val simp_depth_limit_raw: Config.raw
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   107
  val default_mk_sym: Proof.context -> thm -> thm option
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   108
  val simproc_global_i: theory -> string -> term list ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   109
    (Proof.context -> term -> thm option) -> simproc
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   110
  val simproc_global: theory -> string -> string list ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   111
    (Proof.context -> term -> thm option) -> simproc
41227
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
   112
  val simp_trace_depth_limit_raw: Config.raw
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
   113
  val simp_trace_depth_limit_default: int Unsynchronized.ref
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
   114
  val simp_trace_default: bool Unsynchronized.ref
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
   115
  val simp_trace_raw: Config.raw
11e7ee2ca77f tuned signature;
wenzelm
parents: 41226
diff changeset
   116
  val simp_debug_raw: Config.raw
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   117
  val add_prems: thm list -> Proof.context -> Proof.context
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31298
diff changeset
   118
  val debug_bounds: bool Unsynchronized.ref
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   119
  val set_reorient: (Proof.context -> term list -> term -> term -> bool) ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   120
    Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   121
  val set_solvers: solver list -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   122
  val rewrite_cterm: bool * bool * bool ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   123
    (Proof.context -> thm -> thm option) -> Proof.context -> conv
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16380
diff changeset
   124
  val rewrite_term: theory -> thm list -> (term -> term option) list -> term -> term
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   125
  val rewrite_thm: bool * bool * bool ->
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   126
    (Proof.context -> thm -> thm option) -> Proof.context -> thm -> thm
46465
5ba52c337cd0 tuned signature;
wenzelm
parents: 46462
diff changeset
   127
  val generic_rewrite_goal_tac: bool * bool * bool ->
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   128
    (Proof.context -> tactic) -> Proof.context -> int -> tactic
23598
wenzelm
parents: 23584
diff changeset
   129
  val rewrite: bool -> thm list -> conv
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   130
end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   131
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
   132
structure Raw_Simplifier: RAW_SIMPLIFIER =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   133
struct
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   134
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   135
(** datatype simpset **)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   136
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   137
(* congruence rules *)
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   138
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   139
type cong_name = bool * string;
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   140
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   141
fun cong_name (Const (a, _)) = SOME (true, a)
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   142
  | cong_name (Free (a, _)) = SOME (false, a)
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   143
  | cong_name _ = NONE;
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   144
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   145
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   146
(* rewrite rules *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   147
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   148
type rrule =
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   149
 {thm: thm,         (*the rewrite rule*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   150
  name: string,     (*name of theorem from which rewrite rule was extracted*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   151
  lhs: term,        (*the left-hand side*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   152
  elhs: cterm,      (*the etac-contracted lhs*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   153
  extra: bool,      (*extra variables outside of elhs*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   154
  fo: bool,         (*use first-order matching*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   155
  perm: bool};      (*the rewrite rule is permutative*)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   156
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   157
(*
12603
7d2bca103101 tuned tracing markup;
wenzelm
parents: 12285
diff changeset
   158
Remarks:
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   159
  - elhs is used for matching,
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   160
    lhs only for preservation of bound variable names;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   161
  - fo is set iff
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   162
    either elhs is first-order (no Var is applied),
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   163
      in which case fo-matching is complete,
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   164
    or elhs is not a pattern,
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   165
      in which case there is nothing better to do;
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   166
*)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   167
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   168
fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) =
22360
26ead7ed4f4b moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
wenzelm
parents: 22254
diff changeset
   169
  Thm.eq_thm_prop (thm1, thm2);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   170
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   171
(* FIXME: it seems that the conditions on extra variables are too liberal if
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   172
prems are nonempty: does solving the prems really guarantee instantiation of
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   173
all its Vars? Better: a dynamic check each time a rule is applied.
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   174
*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   175
fun rewrite_rule_extra_vars prems elhs erhs =
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   176
  let
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   177
    val elhss = elhs :: prems;
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   178
    val tvars = fold Term.add_tvars elhss [];
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   179
    val vars = fold Term.add_vars elhss [];
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   180
  in
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   181
    erhs |> Term.exists_type (Term.exists_subtype
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   182
      (fn TVar v => not (member (op =) tvars v) | _ => false)) orelse
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   183
    erhs |> Term.exists_subterm
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   184
      (fn Var v => not (member (op =) vars v) | _ => false)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   185
  end;
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   186
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   187
fun rrule_extra_vars elhs thm =
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   188
  rewrite_rule_extra_vars [] (term_of elhs) (Thm.full_prop_of thm);
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   189
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   190
fun mk_rrule2 {thm, name, lhs, elhs, perm} =
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   191
  let
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   192
    val t = term_of elhs;
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   193
    val fo = Pattern.first_order t orelse not (Pattern.pattern t);
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   194
    val extra = rrule_extra_vars elhs thm;
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   195
  in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   196
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   197
(*simple test for looping rewrite rules and stupid orientations*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   198
fun default_reorient ctxt prems lhs rhs =
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   199
  rewrite_rule_extra_vars prems lhs rhs
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   200
    orelse
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   201
  is_Var (head_of lhs)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   202
    orelse
16305
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   203
(* turns t = x around, which causes a headache if x is a local variable -
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   204
   usually it is very useful :-(
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   205
  is_Free rhs andalso not(is_Free lhs) andalso not(Logic.occs(rhs,lhs))
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   206
  andalso not(exists_subterm is_Var lhs)
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   207
    orelse
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   208
*)
16842
wenzelm
parents: 16807
diff changeset
   209
  exists (fn t => Logic.occs (lhs, t)) (rhs :: prems)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   210
    orelse
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   211
  null prems andalso Pattern.matches (Proof_Context.theory_of ctxt) (lhs, rhs)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   212
    (*the condition "null prems" is necessary because conditional rewrites
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   213
      with extra variables in the conditions may terminate although
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   214
      the rhs is an instance of the lhs; example: ?m < ?n ==> f(?n) == f(?m)*)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   215
    orelse
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   216
  is_Const lhs andalso not (is_Const rhs);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   217
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   218
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   219
(* simplification procedures *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   220
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   221
datatype proc =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   222
  Proc of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   223
   {name: string,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   224
    lhs: cterm,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   225
    proc: Proof.context -> cterm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   226
    id: stamp * thm list};
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   227
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   228
fun eq_procid ((s1: stamp, ths1: thm list), (s2, ths2)) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   229
  s1 = s2 andalso eq_list Thm.eq_thm (ths1, ths2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   230
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   231
fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = eq_procid (id1, id2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   232
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   233
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   234
(* solvers *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   235
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   236
datatype solver =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   237
  Solver of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   238
   {name: string,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   239
    solver: Proof.context -> int -> tactic,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   240
    id: stamp};
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   241
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   242
fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()};
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   243
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   244
fun solver_name (Solver {name, ...}) = name;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   245
fun solver ctxt (Solver {solver = tac, ...}) = tac ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   246
fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   247
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   248
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   249
(* simplification sets *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   250
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   251
(*A simpset contains data required during conversion:
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   252
    rules: discrimination net of rewrite rules;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   253
    prems: current premises;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   254
    bounds: maximal index of bound variables already used
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   255
      (for generating new names when rewriting under lambda abstractions);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   256
    depth: simp_depth and exceeded flag;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   257
    congs: association list of congruence rules and
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   258
           a list of `weak' congruence constants.
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   259
           A congruence is `weak' if it avoids normalization of some argument.
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   260
    procs: discrimination net of simplification procedures
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   261
      (functions that prove rewrite rules on the fly);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   262
    mk_rews:
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   263
      mk: turn simplification thms into rewrite rules;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   264
      mk_cong: prepare congruence rules;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   265
      mk_sym: turn == around;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   266
      mk_eq_True: turn P into P == True;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   267
    termless: relation for ordered rewriting;*)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   268
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   269
datatype simpset =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   270
  Simpset of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   271
   {rules: rrule Net.net,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   272
    prems: thm list,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   273
    bounds: int * ((string * typ) * string) list,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   274
    depth: int * bool Unsynchronized.ref} *
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   275
   {congs: (cong_name * thm) list * cong_name list,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   276
    procs: proc Net.net,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   277
    mk_rews:
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   278
     {mk: Proof.context -> thm -> thm list,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   279
      mk_cong: Proof.context -> thm -> thm,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   280
      mk_sym: Proof.context -> thm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   281
      mk_eq_True: Proof.context -> thm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   282
      reorient: Proof.context -> term list -> term -> term -> bool},
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   283
    termless: term * term -> bool,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   284
    subgoal_tac: Proof.context -> int -> tactic,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   285
    loop_tacs: (string * (Proof.context -> int -> tactic)) list,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   286
    solvers: solver list * solver list};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   287
54728
445e7947c6b5 tuned signature;
wenzelm
parents: 54727
diff changeset
   288
fun internal_ss (Simpset (_, ss2)) = ss2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   289
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   290
fun make_ss1 (rules, prems, bounds, depth) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   291
  {rules = rules, prems = prems, bounds = bounds, depth = depth};
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   292
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   293
fun map_ss1 f {rules, prems, bounds, depth} =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   294
  make_ss1 (f (rules, prems, bounds, depth));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   295
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   296
fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   297
  {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   298
    subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   299
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   300
fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} =
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   301
  make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   302
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   303
fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   304
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   305
fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   306
 {simps = Net.entries rules
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   307
    |> map (fn {name, thm, ...} => (name, thm)),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   308
  procs = Net.entries procs
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   309
    |> map (fn Proc {name, lhs, id, ...} => ((name, lhs), id))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   310
    |> partition_eq (eq_snd eq_procid)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   311
    |> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   312
  congs = #1 congs,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   313
  weak_congs = #2 congs,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   314
  loopers = map fst loop_tacs,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   315
  unsafe_solvers = map solver_name (#1 solvers),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   316
  safe_solvers = map solver_name (#2 solvers)};
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   317
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   318
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   319
(* empty *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   320
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   321
fun init_ss mk_rews termless subgoal_tac solvers =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   322
  make_simpset ((Net.empty, [], (0, []), (0, Unsynchronized.ref false)),
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   323
    (([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   324
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   325
fun default_mk_sym _ th = SOME (th RS Drule.symmetric_thm);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   326
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   327
val empty_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   328
  init_ss
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   329
    {mk = fn _ => fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [],
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   330
      mk_cong = K I,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   331
      mk_sym = default_mk_sym,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   332
      mk_eq_True = K (K NONE),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   333
      reorient = default_reorient}
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   334
    Term_Ord.termless (K (K no_tac)) ([], []);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   335
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   336
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   337
(* merge *)  (*NOTE: ignores some fields of 2nd simpset*)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   338
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   339
fun merge_ss (ss1, ss2) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   340
  if pointer_eq (ss1, ss2) then ss1
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   341
  else
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   342
    let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   343
      val Simpset ({rules = rules1, prems = prems1, bounds = bounds1, depth = depth1},
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   344
       {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   345
        loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   346
      val Simpset ({rules = rules2, prems = prems2, bounds = bounds2, depth = depth2},
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   347
       {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   348
        loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   349
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   350
      val rules' = Net.merge eq_rrule (rules1, rules2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   351
      val prems' = Thm.merge_thms (prems1, prems2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   352
      val bounds' = if #1 bounds1 < #1 bounds2 then bounds2 else bounds1;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   353
      val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   354
      val congs' = merge (Thm.eq_thm_prop o pairself #2) (congs1, congs2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   355
      val weak' = merge (op =) (weak1, weak2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   356
      val procs' = Net.merge eq_proc (procs1, procs2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   357
      val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   358
      val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   359
      val solvers' = merge eq_solver (solvers1, solvers2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   360
    in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   361
      make_simpset ((rules', prems', bounds', depth'), ((congs', weak'), procs',
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   362
        mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers')))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   363
    end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   364
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   365
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   366
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   367
(** context data **)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   368
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   369
structure Simpset = Generic_Data
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   370
(
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   371
  type T = simpset;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   372
  val empty = empty_ss;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   373
  val extend = I;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   374
  val merge = merge_ss;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   375
);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   376
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   377
val simpset_of = Simpset.get o Context.Proof;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   378
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   379
fun map_simpset f = Context.proof_map (Simpset.map f);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   380
fun map_simpset1 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (map_ss1 f ss1, ss2));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   381
fun map_simpset2 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (ss1, map_ss2 f ss2));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   382
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   383
fun simpset_map ctxt f ss = ctxt |> map_simpset (K ss) |> f |> Context.Proof |> Simpset.get;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   384
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   385
fun put_simpset ss = map_simpset (fn context_ss =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   386
  let
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   387
    val Simpset ({rules, prems, ...}, ss2) = ss;  (* FIXME prems from context (!?) *)
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   388
    val Simpset ({bounds, depth, ...}, _) = context_ss;
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   389
  in Simpset (make_ss1 (rules, prems, bounds, depth), ss2) end);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   390
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   391
fun global_context thy ss = Proof_Context.init_global thy |> put_simpset ss;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   392
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   393
val empty_simpset = put_simpset empty_ss;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   394
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   395
fun map_theory_simpset f thy =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   396
  let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   397
    val ctxt' = f (Proof_Context.init_global thy);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   398
    val thy' = Proof_Context.theory_of ctxt';
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   399
  in Context.theory_map (Simpset.map (K (simpset_of ctxt'))) thy' end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   400
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   401
fun map_ss f = Context.mapping (map_theory_simpset f) f;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   402
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   403
val clear_simpset =
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   404
  map_simpset (fn Simpset (_, {mk_rews, termless, subgoal_tac, solvers, ...}) =>
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   405
    init_ss mk_rews termless subgoal_tac solvers);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   406
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   407
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   408
(* simp depth *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   409
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   410
val simp_depth_limit_raw = Config.declare "simp_depth_limit" (K (Config.Int 100));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   411
val simp_depth_limit = Config.int simp_depth_limit_raw;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   412
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   413
val simp_trace_depth_limit_default = Unsynchronized.ref 1;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   414
val simp_trace_depth_limit_raw = Config.declare "simp_trace_depth_limit"
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   415
  (fn _ => Config.Int (! simp_trace_depth_limit_default));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   416
val simp_trace_depth_limit = Config.int simp_trace_depth_limit_raw;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   417
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   418
fun trace_depth ctxt msg =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   419
  let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   420
    val Simpset ({depth = (depth, exceeded), ...}, _) = simpset_of ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   421
    val depth_limit = Config.get ctxt simp_trace_depth_limit;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   422
  in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   423
    if depth > depth_limit then
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   424
      if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   425
    else (tracing (enclose "[" "]" (string_of_int depth) ^ msg); exceeded := false)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   426
  end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   427
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   428
fun inc_simp_depth ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   429
  ctxt |> map_simpset1 (fn (rules, prems, bounds, (depth, exceeded)) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   430
    (rules, prems, bounds,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   431
      (depth + 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   432
        if depth = Config.get ctxt simp_trace_depth_limit
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   433
        then Unsynchronized.ref false else exceeded)));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   434
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   435
fun simp_depth ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   436
  let val Simpset ({depth = (depth, _), ...}, _) = simpset_of ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   437
  in depth end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   438
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   439
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   440
(* diagnostics *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   441
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   442
exception SIMPLIFIER of string * thm;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   443
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   444
val simp_debug_raw = Config.declare "simp_debug" (K (Config.Bool false));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   445
val simp_debug = Config.bool simp_debug_raw;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   446
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   447
val simp_trace_default = Unsynchronized.ref false;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   448
val simp_trace_raw = Config.declare "simp_trace" (fn _ => Config.Bool (! simp_trace_default));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   449
val simp_trace = Config.bool simp_trace_raw;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   450
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   451
fun if_enabled ctxt flag f = if Config.get ctxt flag then f ctxt else ();
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   452
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   453
local
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   454
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   455
fun prnt ctxt warn a = if warn then warning a else trace_depth ctxt a;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   456
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   457
fun show_bounds ctxt t =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   458
  let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   459
    val Simpset ({bounds = (_, bs), ...}, _) = simpset_of ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   460
    val names = Term.declare_term_names t Name.context;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   461
    val xs = rev (#1 (fold_map Name.variant (rev (map #2 bs)) names));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   462
    fun subst (((b, T), _), x') = (Free (b, T), Syntax_Trans.mark_bound_abs (x', T));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   463
  in Term.subst_atomic (ListPair.map subst (bs, xs)) t end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   464
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   465
in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   466
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   467
fun print_term ctxt warn a t =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   468
  prnt ctxt warn (a () ^ "\n" ^
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   469
    Syntax.string_of_term ctxt (if Config.get ctxt simp_debug then t else show_bounds ctxt t));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   470
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   471
fun debug ctxt warn a = if_enabled ctxt simp_debug (fn _ => prnt ctxt warn (a ()));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   472
fun trace ctxt warn a = if_enabled ctxt simp_trace (fn _ => prnt ctxt warn (a ()));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   473
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   474
fun debug_term ctxt warn a t = if_enabled ctxt simp_debug (fn _ => print_term ctxt warn a t);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   475
fun trace_term ctxt warn a t = if_enabled ctxt simp_trace (fn _ => print_term ctxt warn a t);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   476
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   477
fun trace_cterm ctxt warn a ct =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   478
  if_enabled ctxt simp_trace (fn _ => print_term ctxt warn a (Thm.term_of ct));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   479
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   480
fun trace_thm ctxt a th =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   481
  if_enabled ctxt simp_trace (fn _ => print_term ctxt false a (Thm.full_prop_of th));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   482
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   483
fun trace_named_thm ctxt a (th, name) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   484
  if_enabled ctxt simp_trace (fn _ =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   485
    print_term ctxt false
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   486
      (fn () => if name = "" then a () else a () ^ " " ^ quote name ^ ":")
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   487
      (Thm.full_prop_of th));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   488
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   489
fun warn_thm ctxt a th = print_term ctxt true a (Thm.full_prop_of th);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   490
fun cond_warn_thm ctxt a th = Context_Position.if_visible ctxt (fn () => warn_thm ctxt a th) ();
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   491
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   492
end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   493
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   494
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   495
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   496
(** simpset operations **)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   497
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   498
(* context *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   499
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   500
fun eq_bound (x: string, (y, _)) = x = y;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   501
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   502
fun add_bound bound =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   503
  map_simpset1 (fn (rules, prems, (count, bounds), depth) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   504
    (rules, prems, (count + 1, bound :: bounds), depth));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   505
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   506
fun prems_of ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   507
  let val Simpset ({prems, ...}, _) = simpset_of ctxt in prems end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   508
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   509
fun add_prems ths =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   510
  map_simpset1 (fn (rules, prems, bounds, depth) => (rules, ths @ prems, bounds, depth));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   511
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   512
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   513
(* maintain simp rules *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   514
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   515
fun del_rrule (rrule as {thm, elhs, ...}) ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   516
  ctxt |> map_simpset1 (fn (rules, prems, bounds, depth) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   517
    (Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, bounds, depth))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   518
  handle Net.DELETE => (cond_warn_thm ctxt (fn () => "Rewrite rule not in simpset:") thm; ctxt);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   519
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   520
fun insert_rrule (rrule as {thm, name, ...}) ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   521
 (trace_named_thm ctxt (fn () => "Adding rewrite rule") (thm, name);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   522
  ctxt |> map_simpset1 (fn (rules, prems, bounds, depth) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   523
    let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   524
      val rrule2 as {elhs, ...} = mk_rrule2 rrule;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   525
      val rules' = Net.insert_term eq_rrule (term_of elhs, rrule2) rules;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   526
    in (rules', prems, bounds, depth) end)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   527
  handle Net.INSERT => (cond_warn_thm ctxt (fn () => "Ignoring duplicate rewrite rule:") thm; ctxt));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   528
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   529
local
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   530
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   531
fun vperm (Var _, Var _) = true
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   532
  | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   533
  | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   534
  | vperm (t, u) = (t = u);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   535
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   536
fun var_perm (t, u) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   537
  vperm (t, u) andalso eq_set (op =) (Term.add_vars t [], Term.add_vars u []);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   538
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   539
in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   540
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   541
fun decomp_simp thm =
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   542
  let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   543
    val prop = Thm.prop_of thm;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   544
    val prems = Logic.strip_imp_prems prop;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   545
    val concl = Drule.strip_imp_concl (Thm.cprop_of thm);
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
   546
    val (lhs, rhs) = Thm.dest_equals concl handle TERM _ =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   547
      raise SIMPLIFIER ("Rewrite rule not a meta-equality", thm);
20579
4dc799edef89 Thm.dest_arg;
wenzelm
parents: 20546
diff changeset
   548
    val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs));
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18573
diff changeset
   549
    val erhs = Envir.eta_contract (term_of rhs);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   550
    val perm =
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   551
      var_perm (term_of elhs, erhs) andalso
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   552
      not (term_of elhs aconv erhs) andalso
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   553
      not (is_Var (term_of elhs));
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   554
  in (prems, term_of lhs, elhs, term_of rhs, perm) end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   555
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   556
end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   557
12783
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   558
fun decomp_simp' thm =
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   559
  let val (_, lhs, _, rhs, _) = decomp_simp thm in
12783
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   560
    if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", thm)
12979
4c76bce4ce39 decomp_simp': use lhs instead of elhs (preserves more bound variable names);
wenzelm
parents: 12783
diff changeset
   561
    else (lhs, rhs)
12783
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   562
  end;
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   563
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   564
fun mk_eq_True ctxt (thm, name) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   565
  let val Simpset (_, {mk_rews = {mk_eq_True, ...}, ...}) = simpset_of ctxt in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   566
    (case mk_eq_True ctxt thm of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   567
      NONE => []
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   568
    | SOME eq_True =>
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   569
        let val (_, lhs, elhs, _, _) = decomp_simp eq_True;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   570
        in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   571
  end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   572
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   573
(*create the rewrite rule and possibly also the eq_True variant,
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   574
  in case there are extra vars on the rhs*)
52082
wenzelm
parents: 52037
diff changeset
   575
fun rrule_eq_True ctxt thm name lhs elhs rhs thm2 =
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   576
  let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   577
    if rewrite_rule_extra_vars [] lhs rhs then
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   578
      mk_eq_True ctxt (thm2, name) @ [rrule]
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   579
    else [rrule]
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   580
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   581
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   582
fun mk_rrule ctxt (thm, name) =
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   583
  let val (prems, lhs, elhs, rhs, perm) = decomp_simp thm in
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   584
    if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   585
    else
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   586
      (*weak test for loops*)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   587
      if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   588
      then mk_eq_True ctxt (thm, name)
52082
wenzelm
parents: 52037
diff changeset
   589
      else rrule_eq_True ctxt thm name lhs elhs rhs thm
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   590
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   591
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   592
fun orient_rrule ctxt (thm, name) =
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   593
  let
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   594
    val (prems, lhs, elhs, rhs, perm) = decomp_simp thm;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   595
    val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = simpset_of ctxt;
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   596
  in
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   597
    if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   598
    else if reorient ctxt prems lhs rhs then
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   599
      if reorient ctxt prems rhs lhs
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   600
      then mk_eq_True ctxt (thm, name)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   601
      else
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   602
        (case mk_sym ctxt thm of
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   603
          NONE => []
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   604
        | SOME thm' =>
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   605
            let val (_, lhs', elhs', rhs', _) = decomp_simp thm'
52082
wenzelm
parents: 52037
diff changeset
   606
            in rrule_eq_True ctxt thm' name lhs' elhs' rhs' thm end)
wenzelm
parents: 52037
diff changeset
   607
    else rrule_eq_True ctxt thm name lhs elhs rhs thm
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   608
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   609
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   610
fun extract_rews (ctxt, thms) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   611
  let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   612
  in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   613
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   614
fun extract_safe_rrules (ctxt, thm) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   615
  maps (orient_rrule ctxt) (extract_rews (ctxt, [thm]));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   616
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   617
20028
b9752164ad92 add/del_simps: warning for inactive simpset (no context);
wenzelm
parents: 19798
diff changeset
   618
(* add/del rules explicitly *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   619
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   620
fun comb_simps comb mk_rrule (ctxt, thms) =
20028
b9752164ad92 add/del_simps: warning for inactive simpset (no context);
wenzelm
parents: 19798
diff changeset
   621
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   622
    val rews = extract_rews (ctxt, thms);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   623
  in fold (fold comb o mk_rrule) rews ctxt end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   624
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   625
fun ctxt addsimps thms =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   626
  comb_simps insert_rrule (mk_rrule ctxt) (ctxt, thms);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   627
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   628
fun ctxt delsimps thms =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   629
  comb_simps del_rrule (map mk_rrule2 o mk_rrule ctxt) (ctxt, thms);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   630
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   631
fun add_simp thm ctxt = ctxt addsimps [thm];
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   632
fun del_simp thm ctxt = ctxt delsimps [thm];
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   633
30318
3d03190d2864 replaced archaic use of rep_ss by Simplifier.mksimps;
wenzelm
parents: 29269
diff changeset
   634
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   635
(* congs *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   636
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   637
local
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   638
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   639
fun is_full_cong_prems [] [] = true
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   640
  | is_full_cong_prems [] _ = false
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   641
  | is_full_cong_prems (p :: prems) varpairs =
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   642
      (case Logic.strip_assums_concl p of
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   643
        Const ("==", _) $ lhs $ rhs =>
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   644
          let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   645
            is_Var x andalso forall is_Bound xs andalso
20972
db0425645cc7 abandoned findrep
haftmann
parents: 20905
diff changeset
   646
            not (has_duplicates (op =) xs) andalso xs = ys andalso
20671
2aa8269a868e member (op =);
wenzelm
parents: 20579
diff changeset
   647
            member (op =) varpairs (x, y) andalso
19303
7da7e96bd74d gen_eq_set, remove (op =);
wenzelm
parents: 19142
diff changeset
   648
            is_full_cong_prems prems (remove (op =) (x, y) varpairs)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   649
          end
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   650
      | _ => false);
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   651
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   652
fun is_full_cong thm =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   653
  let
43597
b4a093e755db tuned signature;
wenzelm
parents: 43596
diff changeset
   654
    val prems = Thm.prems_of thm and concl = Thm.concl_of thm;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   655
    val (lhs, rhs) = Logic.dest_equals concl;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   656
    val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   657
  in
20972
db0425645cc7 abandoned findrep
haftmann
parents: 20905
diff changeset
   658
    f = g andalso not (has_duplicates (op =) (xs @ ys)) andalso length xs = length ys andalso
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   659
    is_full_cong_prems prems (xs ~~ ys)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   660
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   661
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   662
fun mk_cong ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   663
  let val Simpset (_, {mk_rews = {mk_cong = f, ...}, ...}) = simpset_of ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   664
  in f ctxt end;
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45405
diff changeset
   665
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45405
diff changeset
   666
in
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45405
diff changeset
   667
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   668
fun add_eqcong thm ctxt = ctxt |> map_simpset2
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   669
  (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   670
    let
45621
wenzelm
parents: 45620
diff changeset
   671
      val (lhs, _) = Logic.dest_equals (Thm.concl_of thm)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   672
        handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm);
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18573
diff changeset
   673
    (*val lhs = Envir.eta_contract lhs;*)
45621
wenzelm
parents: 45620
diff changeset
   674
      val a = the (cong_name (head_of lhs)) handle Option.Option =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   675
        raise SIMPLIFIER ("Congruence must start with a constant or free variable", thm);
22221
8a8aa6114a89 changed cong alist - now using AList operations instead of overwrite_warn
haftmann
parents: 22008
diff changeset
   676
      val (xs, weak) = congs;
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
   677
      val _ =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   678
        if AList.defined (op =) xs a then
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   679
          Context_Position.if_visible ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   680
            warning ("Overwriting congruence rule for " ^ quote (#2 a))
22221
8a8aa6114a89 changed cong alist - now using AList operations instead of overwrite_warn
haftmann
parents: 22008
diff changeset
   681
        else ();
30908
7ccf4a3d764c replace type cong = {thm : thm, lhs : term} by plain thm -- the other component has been unused for a long time.
krauss
parents: 30552
diff changeset
   682
      val xs' = AList.update (op =) (a, thm) xs;
22221
8a8aa6114a89 changed cong alist - now using AList operations instead of overwrite_warn
haftmann
parents: 22008
diff changeset
   683
      val weak' = if is_full_cong thm then weak else a :: weak;
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   684
    in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   685
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   686
fun del_eqcong thm ctxt = ctxt |> map_simpset2
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   687
  (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   688
    let
45621
wenzelm
parents: 45620
diff changeset
   689
      val (lhs, _) = Logic.dest_equals (Thm.concl_of thm)
wenzelm
parents: 45620
diff changeset
   690
        handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", thm);
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18573
diff changeset
   691
    (*val lhs = Envir.eta_contract lhs;*)
20057
058e913bac71 tuned exception handling;
wenzelm
parents: 20028
diff changeset
   692
      val a = the (cong_name (head_of lhs)) handle Option.Option =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   693
        raise SIMPLIFIER ("Congruence must start with a constant", thm);
22221
8a8aa6114a89 changed cong alist - now using AList operations instead of overwrite_warn
haftmann
parents: 22008
diff changeset
   694
      val (xs, _) = congs;
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   695
      val xs' = filter_out (fn (x : cong_name, _) => x = a) xs;
30908
7ccf4a3d764c replace type cong = {thm : thm, lhs : term} by plain thm -- the other component has been unused for a long time.
krauss
parents: 30552
diff changeset
   696
      val weak' = xs' |> map_filter (fn (a, thm) =>
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
   697
        if is_full_cong thm then NONE else SOME a);
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   698
    in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   699
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   700
fun add_cong thm ctxt = add_eqcong (mk_cong ctxt thm) ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   701
fun del_cong thm ctxt = del_eqcong (mk_cong ctxt thm) ctxt;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   702
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   703
end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   704
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   705
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   706
(* simprocs *)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   707
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   708
datatype simproc =
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   709
  Simproc of
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   710
    {name: string,
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   711
     lhss: cterm list,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   712
     proc: morphism -> Proof.context -> cterm -> thm option,
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   713
     id: stamp * thm list};
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   714
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   715
fun eq_simproc (Simproc {id = id1, ...}, Simproc {id = id2, ...}) = eq_procid (id1, id2);
22008
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   716
45290
f599ac41e7f5 tuned signature -- refined terminology;
wenzelm
parents: 44058
diff changeset
   717
fun transform_simproc phi (Simproc {name, lhss, proc, id = (s, ths)}) =
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   718
  Simproc
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   719
   {name = name,
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   720
    lhss = map (Morphism.cterm phi) lhss,
22669
62857ad97cca Morphism.transform/form;
wenzelm
parents: 22360
diff changeset
   721
    proc = Morphism.transform phi proc,
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   722
    id = (s, Morphism.fact phi ths)};
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   723
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   724
fun make_simproc {name, lhss, proc, identifier} =
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   725
  Simproc {name = name, lhss = lhss, proc = proc, id = (stamp (), identifier)};
22008
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   726
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   727
fun mk_simproc name lhss proc =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   728
  make_simproc {name = name, lhss = lhss, proc = fn _ => fn ctxt => fn ct =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   729
    proc ctxt (Thm.term_of ct), identifier = []};
22008
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   730
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35408
diff changeset
   731
(* FIXME avoid global thy and Logic.varify_global *)
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 36944
diff changeset
   732
fun simproc_global_i thy name = mk_simproc name o map (Thm.cterm_of thy o Logic.varify_global);
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 36944
diff changeset
   733
fun simproc_global thy name = simproc_global_i thy name o map (Syntax.read_term_global thy);
22008
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   734
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   735
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   736
local
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   737
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   738
fun add_proc (proc as Proc {name, lhs, ...}) ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   739
 (trace_cterm ctxt false (fn () => "Adding simplification procedure " ^ quote name ^ " for") lhs;
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   740
  ctxt |> map_simpset2
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   741
    (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   742
      (congs, Net.insert_term eq_proc (term_of lhs, proc) procs,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   743
        mk_rews, termless, subgoal_tac, loop_tacs, solvers))
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   744
  handle Net.INSERT =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   745
    (Context_Position.if_visible ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   746
      warning ("Ignoring duplicate simplification procedure " ^ quote name); ctxt));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   747
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   748
fun del_proc (proc as Proc {name, lhs, ...}) ctxt =
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   749
  ctxt |> map_simpset2
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   750
    (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   751
      (congs, Net.delete_term eq_proc (term_of lhs, proc) procs,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   752
        mk_rews, termless, subgoal_tac, loop_tacs, solvers))
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   753
  handle Net.DELETE =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   754
    (Context_Position.if_visible ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   755
      warning ("Simplification procedure " ^ quote name ^ " not in simpset"); ctxt);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   756
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   757
fun prep_procs (Simproc {name, lhss, proc, id}) =
22669
62857ad97cca Morphism.transform/form;
wenzelm
parents: 22360
diff changeset
   758
  lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, id = id});
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   759
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   760
in
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   761
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   762
fun ctxt addsimprocs ps = fold (fold add_proc o prep_procs) ps ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   763
fun ctxt delsimprocs ps = fold (fold del_proc o prep_procs) ps ctxt;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   764
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   765
end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   766
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   767
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   768
(* mk_rews *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   769
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   770
local
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   771
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   772
fun map_mk_rews f =
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   773
  map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   774
    let
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   775
      val {mk, mk_cong, mk_sym, mk_eq_True, reorient} = mk_rews;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   776
      val (mk', mk_cong', mk_sym', mk_eq_True', reorient') =
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   777
        f (mk, mk_cong, mk_sym, mk_eq_True, reorient);
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   778
      val mk_rews' = {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True',
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   779
        reorient = reorient'};
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   780
    in (congs, procs, mk_rews', termless, subgoal_tac, loop_tacs, solvers) end);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   781
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   782
in
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   783
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   784
fun mksimps ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   785
  let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   786
  in mk ctxt end;
30318
3d03190d2864 replaced archaic use of rep_ss by Simplifier.mksimps;
wenzelm
parents: 29269
diff changeset
   787
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   788
fun set_mksimps mk = map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True, reorient) =>
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   789
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   790
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   791
fun set_mkcong mk_cong = map_mk_rews (fn (mk, _, mk_sym, mk_eq_True, reorient) =>
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   792
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   793
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   794
fun set_mksym mk_sym = map_mk_rews (fn (mk, mk_cong, _, mk_eq_True, reorient) =>
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   795
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   796
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   797
fun set_mkeqTrue mk_eq_True = map_mk_rews (fn (mk, mk_cong, mk_sym, _, reorient) =>
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   798
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   799
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   800
fun set_reorient reorient = map_mk_rews (fn (mk, mk_cong, mk_sym, mk_eq_True, _) =>
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   801
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   802
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   803
end;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   804
14242
ec70653a02bf Added access to the mk_rews field (and friends).
skalberg
parents: 14040
diff changeset
   805
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   806
(* termless *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   807
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   808
fun set_termless termless =
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   809
  map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) =>
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   810
   (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   811
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   812
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   813
(* tactics *)
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   814
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   815
fun set_subgoaler subgoal_tac =
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   816
  map_simpset2 (fn (congs, procs, mk_rews, termless, _, loop_tacs, solvers) =>
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   817
   (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   818
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
   819
fun ctxt setloop tac = ctxt |>
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   820
  map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, _, solvers) =>
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   821
   (congs, procs, mk_rews, termless, subgoal_tac, [("", tac)], solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   822
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
   823
fun ctxt addloop (name, tac) = ctxt |>
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   824
  map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   825
    (congs, procs, mk_rews, termless, subgoal_tac,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   826
     AList.update (op =) (name, tac) loop_tacs, solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   827
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   828
fun ctxt delloop name = ctxt |>
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   829
  map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>
21286
b5e7b80caa6a introduces canonical AList functions for loop_tacs
haftmann
parents: 20972
diff changeset
   830
    (congs, procs, mk_rews, termless, subgoal_tac,
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
   831
     (if AList.defined (op =) loop_tacs name then ()
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   832
      else
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   833
        Context_Position.if_visible ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   834
          warning ("No such looper in simpset: " ^ quote name);
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   835
        AList.delete (op =) name loop_tacs), solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   836
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   837
fun ctxt setSSolver solver = ctxt |> map_simpset2
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   838
  (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, _)) =>
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   839
    (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, [solver])));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   840
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   841
fun ctxt addSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   842
  subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless,
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   843
    subgoal_tac, loop_tacs, (unsafe_solvers, insert eq_solver solver solvers)));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   844
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   845
fun ctxt setSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   846
  subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, termless,
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   847
    subgoal_tac, loop_tacs, ([solver], solvers)));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   848
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   849
fun ctxt addSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   850
  subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless,
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   851
    subgoal_tac, loop_tacs, (insert eq_solver solver unsafe_solvers, solvers)));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   852
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   853
fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, termless,
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   854
  subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, termless,
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   855
  subgoal_tac, loop_tacs, (solvers, solvers)));
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   856
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   857
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   858
(* trace operations *)
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   859
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   860
type trace_ops =
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   861
 {trace_invoke: {depth: int, term: term} -> Proof.context -> Proof.context,
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   862
  trace_apply: {unconditional: bool, term: term, thm: thm, name: string} ->
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   863
    Proof.context -> (Proof.context -> (thm * term) option) -> (thm * term) option};
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   864
54731
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   865
structure Trace_Ops = Theory_Data
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   866
(
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   867
  type T = trace_ops;
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   868
  val empty: T =
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   869
   {trace_invoke = fn _ => fn ctxt => ctxt,
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   870
    trace_apply = fn _ => fn ctxt => fn cont => cont ctxt};
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   871
  val extend = I;
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   872
  fun merge (trace_ops, _) = trace_ops;
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   873
);
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   874
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   875
val set_trace_ops = Trace_Ops.put;
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   876
384ac33802b0 clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
wenzelm
parents: 54729
diff changeset
   877
val trace_ops = Trace_Ops.get o Proof_Context.theory_of;
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   878
fun trace_invoke args ctxt = #trace_invoke (trace_ops ctxt) args ctxt;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   879
fun trace_apply args ctxt = #trace_apply (trace_ops ctxt) args ctxt;
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   880
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   881
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   882
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   883
(** rewriting **)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   884
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   885
(*
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   886
  Uses conversions, see:
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   887
    L C Paulson, A higher-order implementation of rewriting,
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   888
    Science of Computer Programming 3 (1983), pages 119-149.
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   889
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   890
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   891
fun check_conv ctxt msg thm thm' =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   892
  let
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
   893
    val thm'' = Thm.transitive thm thm' handle THM _ =>
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
   894
     Thm.transitive thm (Thm.transitive
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
   895
       (Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm')
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   896
  in if msg then trace_thm ctxt (fn () => "SUCCEEDED") thm' else (); SOME thm'' end
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   897
  handle THM _ =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   898
    let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   899
      val _ $ _ $ prop0 = Thm.prop_of thm;
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   900
    in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   901
      trace_thm ctxt (fn () => "Proved wrong thm (Check subgoaler?)") thm';
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   902
      trace_term ctxt false (fn () => "Should have proved:") prop0;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
   903
      NONE
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   904
    end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   905
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   906
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   907
(* mk_procrule *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   908
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   909
fun mk_procrule ctxt thm =
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   910
  let val (prems, lhs, elhs, rhs, _) = decomp_simp thm in
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   911
    if rewrite_rule_extra_vars prems lhs rhs
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   912
    then (cond_warn_thm ctxt (fn () => "Extra vars on rhs:") thm; [])
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   913
    else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}]
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   914
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   915
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   916
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   917
(* rewritec: conversion to apply the meta simpset to a term *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   918
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   919
(*Since the rewriting strategy is bottom-up, we avoid re-normalizing already
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   920
  normalized terms by carrying around the rhs of the rewrite rule just
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   921
  applied. This is called the `skeleton'. It is decomposed in parallel
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   922
  with the term. Once a Var is encountered, the corresponding term is
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   923
  already in normal form.
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   924
  skel0 is a dummy skeleton that is to enforce complete normalization.*)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   925
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   926
val skel0 = Bound 0;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   927
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   928
(*Use rhs as skeleton only if the lhs does not contain unnormalized bits.
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   929
  The latter may happen iff there are weak congruence rules for constants
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   930
  in the lhs.*)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   931
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   932
fun uncond_skel ((_, weak), (lhs, rhs)) =
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   933
  if null weak then rhs  (*optimization*)
51591
e4aeb102ad70 amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
wenzelm
parents: 51590
diff changeset
   934
  else if exists_subterm
e4aeb102ad70 amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
wenzelm
parents: 51590
diff changeset
   935
    (fn Const (a, _) => member (op =) weak (true, a)
e4aeb102ad70 amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
wenzelm
parents: 51590
diff changeset
   936
      | Free (a, _) => member (op =) weak (false, a)
e4aeb102ad70 amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
wenzelm
parents: 51590
diff changeset
   937
      | _ => false) lhs then skel0
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   938
  else rhs;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   939
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   940
(*Behaves like unconditional rule if rhs does not contain vars not in the lhs.
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   941
  Otherwise those vars may become instantiated with unnormalized terms
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   942
  while the premises are solved.*)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   943
32797
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
   944
fun cond_skel (args as (_, (lhs, rhs))) =
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   945
  if subset (op =) (Term.add_vars rhs [], Term.add_vars lhs []) then uncond_skel args
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   946
  else skel0;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   947
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   948
(*
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   949
  Rewriting -- we try in order:
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   950
    (1) beta reduction
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   951
    (2) unconditional rewrite rules
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   952
    (3) conditional rewrite rules
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   953
    (4) simplification procedures
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   954
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   955
  IMPORTANT: rewrite rules must not introduce new Vars or TVars!
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   956
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   957
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   958
fun rewritec (prover, maxt) ctxt t =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   959
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   960
    val Simpset ({rules, ...}, {congs, procs, termless, ...}) = simpset_of ctxt;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   961
    val eta_thm = Thm.eta_conversion t;
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
   962
    val eta_t' = Thm.rhs_of eta_thm;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   963
    val eta_t = term_of eta_t';
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   964
    fun rew {thm, name, lhs, elhs, extra, fo, perm} =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   965
      let
32797
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
   966
        val prop = Thm.prop_of thm;
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   967
        val (rthm, elhs') =
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   968
          if maxt = ~1 orelse not extra then (thm, elhs)
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
   969
          else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs);
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
   970
        val insts =
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
   971
          if fo then Thm.first_order_match (elhs', eta_t')
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
   972
          else Thm.match (elhs', eta_t');
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   973
        val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm);
14643
wenzelm
parents: 14330
diff changeset
   974
        val prop' = Thm.prop_of thm';
21576
8c11b1ce2f05 simplified Logic.count_prems;
wenzelm
parents: 21565
diff changeset
   975
        val unconditional = (Logic.count_prems prop' = 0);
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   976
        val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop');
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   977
        val trace_args = {unconditional = unconditional, term = eta_t, thm = thm', name = name};
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   978
      in
11295
66925f23ac7f improved tracing of permutative rules.
nipkow
parents: 11291
diff changeset
   979
        if perm andalso not (termless (rhs', lhs'))
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   980
        then
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   981
         (trace_named_thm ctxt (fn () => "Cannot apply permutative rewrite rule") (thm, name);
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   982
          trace_thm ctxt (fn () => "Term does not become smaller:") thm';
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   983
          NONE)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   984
        else
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   985
         (trace_named_thm ctxt (fn () => "Applying instance of rewrite rule") (thm, name);
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   986
          if unconditional
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   987
          then
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   988
           (trace_thm ctxt (fn () => "Rewriting:") thm';
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   989
            trace_apply trace_args ctxt (fn ctxt' =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   990
              let
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   991
                val lr = Logic.dest_equals prop;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   992
                val SOME thm'' = check_conv ctxt' false eta_thm thm';
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   993
              in SOME (thm'', uncond_skel (congs, lr)) end))
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   994
          else
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   995
           (trace_thm ctxt (fn () => "Trying to rewrite:") thm';
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   996
            if simp_depth ctxt > Config.get ctxt simp_depth_limit
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   997
            then
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
   998
              let
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
   999
                val s = "simp_depth_limit exceeded - giving up";
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1000
                val _ = trace ctxt false (fn () => s);
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1001
                val _ = Context_Position.if_visible ctxt warning s;
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1002
              in NONE end
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1003
            else
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1004
              trace_apply trace_args ctxt (fn ctxt' =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1005
                (case prover ctxt' thm' of
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1006
                  NONE => (trace_thm ctxt' (fn () => "FAILED") thm'; NONE)
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1007
                | SOME thm2 =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1008
                    (case check_conv ctxt' true eta_thm thm2 of
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1009
                      NONE => NONE
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1010
                    | SOME thm2' =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1011
                        let
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1012
                          val concl = Logic.strip_imp_concl prop;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1013
                          val lr = Logic.dest_equals concl;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1014
                        in SOME (thm2', cond_skel (congs, lr)) end)))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1015
      end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1016
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1017
    fun rews [] = NONE
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1018
      | rews (rrule :: rrules) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1019
          let val opt = rew rrule handle Pattern.MATCH => NONE
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1020
          in (case opt of NONE => rews rrules | some => some) end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1021
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1022
    fun sort_rrules rrs =
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1023
      let
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1024
        fun is_simple ({thm, ...}: rrule) =
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1025
          (case Thm.prop_of thm of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1026
            Const ("==", _) $ _ $ _ => true
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1027
          | _ => false);
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1028
        fun sort [] (re1, re2) = re1 @ re2
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1029
          | sort (rr :: rrs) (re1, re2) =
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1030
              if is_simple rr
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1031
              then sort rrs (rr :: re1, re2)
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1032
              else sort rrs (re1, rr :: re2);
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1033
      in sort rrs ([], []) end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1034
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1035
    fun proc_rews [] = NONE
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1036
      | proc_rews (Proc {name, proc, lhs, ...} :: ps) =
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1037
          if Pattern.matches (Proof_Context.theory_of ctxt) (Thm.term_of lhs, Thm.term_of t) then
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1038
            (debug_term ctxt false (fn () => "Trying procedure " ^ quote name ^ " on:") eta_t;
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1039
             (case proc ctxt eta_t' of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1040
               NONE => (debug ctxt false (fn () => "FAILED"); proc_rews ps)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1041
             | SOME raw_thm =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1042
                 (trace_thm ctxt (fn () => "Procedure " ^ quote name ^ " produced rewrite rule:")
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1043
                   raw_thm;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1044
                  (case rews (mk_procrule ctxt raw_thm) of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1045
                    NONE => (trace_cterm ctxt true (fn () => "IGNORED result of simproc " ^ quote name ^
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1046
                      " -- does not match") t; proc_rews ps)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1047
                  | some => some))))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1048
          else proc_rews ps;
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1049
  in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1050
    (case eta_t of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1051
      Abs _ $ _ => SOME (Thm.transitive eta_thm (Thm.beta_conversion false eta_t'), skel0)
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1052
    | _ =>
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1053
      (case rews (sort_rrules (Net.match_term rules eta_t)) of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1054
        NONE => proc_rews (Net.match_term procs eta_t)
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1055
      | some => some))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1056
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1057
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1058
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1059
(* conversion to apply a congruence rule to a term *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1060
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1061
fun congc prover ctxt maxt cong t =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1062
  let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1063
    val rthm = Thm.incr_indexes (maxt + 1) cong;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1064
    val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (cprop_of rthm)));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1065
    val insts = Thm.match (rlhs, t)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1066
    (* Thm.match can raise Pattern.MATCH;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1067
       is handled when congc is called *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1068
    val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1069
    val _ = trace_thm ctxt (fn () => "Applying congruence rule:") thm';
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1070
    fun err (msg, thm) = (trace_thm ctxt (fn () => msg) thm; NONE);
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1071
  in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1072
    (case prover thm' of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1073
      NONE => err ("Congruence proof failed.  Could not prove", thm')
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1074
    | SOME thm2 =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1075
        (case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1076
          NONE => err ("Congruence proof failed.  Should not have proved", thm2)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1077
        | SOME thm2' =>
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1078
            if op aconv (pairself term_of (Thm.dest_equals (cprop_of thm2')))
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1079
            then NONE else SOME thm2'))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1080
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1081
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1082
val (cA, (cB, cC)) =
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1083
  apsnd Thm.dest_equals (Thm.dest_implies (hd (cprems_of Drule.imp_cong)));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1084
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1085
fun transitive1 NONE NONE = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1086
  | transitive1 (SOME thm1) NONE = SOME thm1
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1087
  | transitive1 NONE (SOME thm2) = SOME thm2
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1088
  | transitive1 (SOME thm1) (SOME thm2) = SOME (Thm.transitive thm1 thm2);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1089
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1090
fun transitive2 thm = transitive1 (SOME thm);
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1091
fun transitive3 thm = transitive1 thm o SOME;
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1092
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1093
fun bottomc ((simprem, useprem, mutsimp), prover, maxidx) =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1094
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1095
    fun botc skel ctxt t =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1096
      if is_Var skel then NONE
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1097
      else
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1098
        (case subc skel ctxt t of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1099
           some as SOME thm1 =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1100
             (case rewritec (prover, maxidx) ctxt (Thm.rhs_of thm1) of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1101
                SOME (thm2, skel2) =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1102
                  transitive2 (Thm.transitive thm1 thm2)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1103
                    (botc skel2 ctxt (Thm.rhs_of thm2))
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1104
              | NONE => some)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1105
         | NONE =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1106
             (case rewritec (prover, maxidx) ctxt t of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1107
                SOME (thm2, skel2) => transitive2 thm2
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1108
                  (botc skel2 ctxt (Thm.rhs_of thm2))
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1109
              | NONE => NONE))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1110
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1111
    and try_botc ctxt t =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1112
      (case botc skel0 ctxt t of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1113
        SOME trec1 => trec1
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1114
      | NONE => Thm.reflexive t)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1115
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1116
    and subc skel ctxt t0 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1117
        let val Simpset ({bounds, ...}, {congs, ...}) = simpset_of ctxt in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1118
          (case term_of t0 of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1119
              Abs (a, T, _) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1120
                let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1121
                    val b = Name.bound (#1 bounds);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1122
                    val (v, t') = Thm.dest_abs (SOME b) t0;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1123
                    val b' = #1 (Term.dest_Free (Thm.term_of v));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1124
                    val _ =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1125
                      if b <> b' then
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1126
                        warning ("Simplifier: renamed bound variable " ^
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1127
                          quote b ^ " to " ^ quote b' ^ Position.here (Position.thread_data ()))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1128
                      else ();
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1129
                    val ctxt' = add_bound ((b', T), a) ctxt;
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1130
                    val skel' = (case skel of Abs (_, _, sk) => sk | _ => skel0);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1131
                in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1132
                  (case botc skel' ctxt' t' of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1133
                    SOME thm => SOME (Thm.abstract_rule a v thm)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1134
                  | NONE => NONE)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1135
                end
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1136
            | t $ _ =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1137
              (case t of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1138
                Const ("==>", _) $ _  => impc t0 ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1139
              | Abs _ =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1140
                  let val thm = Thm.beta_conversion false t0
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1141
                  in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1142
                    (case subc skel0 ctxt (Thm.rhs_of thm) of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1143
                      NONE => SOME thm
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1144
                    | SOME thm' => SOME (Thm.transitive thm thm'))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1145
                  end
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1146
              | _  =>
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1147
                  let
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1148
                    fun appc () =
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1149
                      let
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1150
                        val (tskel, uskel) =
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1151
                          (case skel of
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1152
                            tskel $ uskel => (tskel, uskel)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1153
                          | _ => (skel0, skel0));
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1154
                        val (ct, cu) = Thm.dest_comb t0;
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1155
                      in
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1156
                        (case botc tskel ctxt ct of
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1157
                          SOME thm1 =>
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1158
                            (case botc uskel ctxt cu of
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1159
                              SOME thm2 => SOME (Thm.combination thm1 thm2)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1160
                            | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu)))
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1161
                        | NONE =>
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1162
                            (case botc uskel ctxt cu of
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1163
                              SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1164
                            | NONE => NONE))
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1165
                      end;
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1166
                    val (h, ts) = strip_comb t;
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1167
                  in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1168
                    (case cong_name h of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1169
                      SOME a =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1170
                        (case AList.lookup (op =) (fst congs) a of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1171
                           NONE => appc ()
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1172
                        | SOME cong =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1173
     (*post processing: some partial applications h t1 ... tj, j <= length ts,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1174
       may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1175
                           (let
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1176
                              val thm = congc (prover ctxt) ctxt maxidx cong t0;
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1177
                              val t = the_default t0 (Option.map Thm.rhs_of thm);
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1178
                              val (cl, cr) = Thm.dest_comb t
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1179
                              val dVar = Var(("", 0), dummyT)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1180
                              val skel =
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1181
                                list_comb (h, replicate (length ts) dVar)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1182
                            in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1183
                              (case botc skel ctxt cl of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1184
                                NONE => thm
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1185
                              | SOME thm' =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1186
                                  transitive3 thm (Thm.combination thm' (Thm.reflexive cr)))
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1187
                            end handle Pattern.MATCH => appc ()))
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1188
                     | _ => appc ())
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1189
                  end)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1190
            | _ => NONE)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1191
        end
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1192
    and impc ct ctxt =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1193
      if mutsimp then mut_impc0 [] ct [] [] ctxt
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1194
      else nonmut_impc ct ctxt
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1195
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1196
    and rules_of_prem ctxt prem =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1197
      if maxidx_of_term (term_of prem) <> ~1
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1198
      then (trace_cterm ctxt true
22254
420625970f31 trace/debug: avoid eager string concatenation;
wenzelm
parents: 22234
diff changeset
  1199
        (fn () => "Cannot add premise as rewrite rule because it contains (type) unknowns:")
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1200
        prem; ([], NONE))
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1201
      else
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  1202
        let val asm = Thm.assume prem
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1203
        in (extract_safe_rrules (ctxt, asm), SOME asm) end
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1204
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1205
    and add_rrules (rrss, asms) ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1206
      (fold o fold) insert_rrule rrss ctxt |> add_prems (map_filter I asms)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1207
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22902
diff changeset
  1208
    and disch r prem eq =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1209
      let
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1210
        val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq);
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1211
        val eq' =
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1212
          Thm.implies_elim
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1213
            (Thm.instantiate ([], [(cA, prem), (cB, lhs), (cC, rhs)]) Drule.imp_cong)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1214
            (Thm.implies_intr prem eq);
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1215
      in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1216
        if not r then eq'
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1217
        else
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1218
          let
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1219
            val (prem', concl) = Thm.dest_implies lhs;
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1220
            val (prem'', _) = Thm.dest_implies rhs;
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1221
          in
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1222
            Thm.transitive
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1223
              (Thm.transitive
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1224
                (Thm.instantiate ([], [(cA, prem'), (cB, prem), (cC, concl)]) Drule.swap_prems_eq)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1225
                eq')
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1226
              (Thm.instantiate ([], [(cA, prem), (cB, prem''), (cC, concl)]) Drule.swap_prems_eq)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1227
          end
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1228
      end
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1229
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1230
    and rebuild [] _ _ _ _ eq = eq
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1231
      | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ctxt eq =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1232
          let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1233
            val ctxt' = add_rrules (rev rrss, rev asms) ctxt;
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1234
            val concl' =
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1235
              Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq));
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1236
            val dprem = Option.map (disch false prem);
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1237
          in
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1238
            (case rewritec (prover, maxidx) ctxt' concl' of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1239
              NONE => rebuild prems concl' rrss asms ctxt (dprem eq)
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1240
            | SOME (eq', _) =>
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1241
                transitive2 (fold (disch false) prems (the (transitive3 (dprem eq) eq')))
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1242
                  (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ctxt))
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1243
          end
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1244
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1245
    and mut_impc0 prems concl rrss asms ctxt =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1246
      let
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1247
        val prems' = strip_imp_prems concl;
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1248
        val (rrss', asms') = split_list (map (rules_of_prem ctxt) prems');
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1249
      in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1250
        mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss')
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1251
          (asms @ asms') [] [] [] [] ctxt ~1 ~1
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1252
      end
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1253
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1254
    and mut_impc [] concl [] [] prems' rrss' asms' eqns ctxt changed k =
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33039
diff changeset
  1255
        transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33039
diff changeset
  1256
            (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE)
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1257
          (if changed > 0 then
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1258
             mut_impc (rev prems') concl (rev rrss') (rev asms')
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1259
               [] [] [] [] ctxt ~1 changed
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1260
           else rebuild prems' concl rrss' asms' ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1261
             (botc skel0 (add_rrules (rev rrss', rev asms') ctxt) concl))
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1262
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1263
      | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1264
          prems' rrss' asms' eqns ctxt changed k =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1265
        (case (if k = 0 then NONE else botc skel0 (add_rrules
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1266
          (rev rrss' @ rrss, rev asms' @ asms) ctxt) prem) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1267
            NONE => mut_impc prems concl rrss asms (prem :: prems')
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1268
              (rrs :: rrss') (asm :: asms') (NONE :: eqns) ctxt changed
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1269
              (if k = 0 then 0 else k - 1)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1270
        | SOME eqn =>
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1271
            let
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1272
              val prem' = Thm.rhs_of eqn;
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1273
              val tprems = map term_of prems;
33029
2fefe039edf1 uniform use of Integer.min/max;
wenzelm
parents: 32797
diff changeset
  1274
              val i = 1 + fold Integer.max (map (fn p =>
44058
ae85c5d64913 misc tuning -- eliminated old-fashioned rep_thm;
wenzelm
parents: 43597
diff changeset
  1275
                find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1276
              val (rrs', asm') = rules_of_prem ctxt prem';
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1277
            in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1278
              mut_impc prems concl rrss asms (prem' :: prems')
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1279
                (rrs' :: rrss') (asm' :: asms')
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1280
                (SOME (fold_rev (disch true)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1281
                  (take i prems)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1282
                  (Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1283
                    (drop i prems, concl))))) :: eqns)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1284
                ctxt (length prems') ~1
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1285
            end)
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1286
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1287
    (*legacy code -- only for backwards compatibility*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1288
    and nonmut_impc ct ctxt =
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1289
      let
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1290
        val (prem, conc) = Thm.dest_implies ct;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1291
        val thm1 = if simprem then botc skel0 ctxt prem else NONE;
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1292
        val prem1 = the_default prem (Option.map Thm.rhs_of thm1);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1293
        val ctxt1 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1294
          if not useprem then ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1295
          else add_rrules (apsnd single (apfst single (rules_of_prem ctxt prem1))) ctxt
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1296
      in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1297
        (case botc skel0 ctxt1 conc of
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1298
          NONE =>
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1299
            (case thm1 of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1300
              NONE => NONE
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1301
            | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc)))
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1302
        | SOME thm2 =>
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1303
            let val thm2' = disch false prem1 thm2 in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1304
              (case thm1 of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1305
                NONE => SOME thm2'
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1306
              | SOME thm1' =>
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  1307
                 SOME (Thm.transitive (Drule.imp_cong_rule thm1' (Thm.reflexive conc)) thm2'))
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1308
            end)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1309
      end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1310
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1311
  in try_botc end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1312
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1313
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1314
(* Meta-rewriting: rewrites t to u and returns the theorem t==u *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1315
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1316
(*
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1317
  Parameters:
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1318
    mode = (simplify A,
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1319
            use A in simplifying B,
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1320
            use prems of B (if B is again a meta-impl.) to simplify A)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1321
           when simplifying A ==> B
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1322
    prover: how to solve premises in conditional rewrites and congruences
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1323
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1324
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31298
diff changeset
  1325
val debug_bounds = Unsynchronized.ref false;
17705
a5d146aca659 removed revert_bound;
wenzelm
parents: 17614
diff changeset
  1326
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1327
fun check_bounds ctxt ct =
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21708
diff changeset
  1328
  if ! debug_bounds then
279b129498b6 removed conditional combinator;
wenzelm
parents: 21708
diff changeset
  1329
    let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1330
      val Simpset ({bounds = (_, bounds), ...}, _) = simpset_of ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1331
      val bs =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1332
        fold_aterms
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1333
          (fn Free (x, _) =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1334
            if Name.is_bound x andalso not (AList.defined eq_bound bounds x)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1335
            then insert (op =) x else I
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1336
          | _ => I) (term_of ct) [];
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21708
diff changeset
  1337
    in
279b129498b6 removed conditional combinator;
wenzelm
parents: 21708
diff changeset
  1338
      if null bs then ()
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1339
      else
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1340
        print_term ctxt true
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1341
          (fn () => "Simplifier: term contains loose bounds: " ^ commas_quote bs)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1342
          (Thm.term_of ct)
21962
279b129498b6 removed conditional combinator;
wenzelm
parents: 21708
diff changeset
  1343
    end
279b129498b6 removed conditional combinator;
wenzelm
parents: 21708
diff changeset
  1344
  else ();
17614
37ee526db497 added mk_solver';
wenzelm
parents: 17496
diff changeset
  1345
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1346
fun rewrite_cterm mode prover raw_ctxt raw_ct =
17882
b6e44fc46cf0 added set/addloop' for simpset dependent loopers;
wenzelm
parents: 17756
diff changeset
  1347
  let
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1348
    val thy = Proof_Context.theory_of raw_ctxt;
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1349
20260
990dbc007ca6 Thm.adjust_maxidx;
wenzelm
parents: 20228
diff changeset
  1350
    val ct = Thm.adjust_maxidx_cterm ~1 raw_ct;
32797
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
  1351
    val {maxidx, ...} = Thm.rep_cterm ct;
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1352
    val _ =
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1353
      Theory.subthy (theory_of_cterm ct, thy) orelse
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1354
        raise CTERM ("rewrite_cterm: bad background theory", [ct]);
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1355
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1356
    val ctxt =
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1357
      raw_ctxt
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1358
      |> Context_Position.set_visible false
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1359
      |> inc_simp_depth
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1360
      |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, term = Thm.term_of ct} ctxt);
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1361
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1362
    val _ = trace_cterm ctxt false (fn () => "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:") ct;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1363
    val _ = check_bounds ctxt ct;
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1364
  in bottomc (mode, Option.map Drule.flexflex_unique oo prover, maxidx) ctxt ct end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1365
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1366
val simple_prover =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1367
  SINGLE o (fn ctxt => ALLGOALS (resolve_tac (prems_of ctxt)));
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1368
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1369
fun rewrite _ [] ct = Thm.reflexive ct
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1370
  | rewrite full thms ct =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1371
      rewrite_cterm (full, false, false) simple_prover
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1372
        (global_context (Thm.theory_of_cterm ct) empty_ss addsimps thms) ct;
11672
8e75b78f33f3 full_rewrite_cterm_aux (see also tactic.ML);
wenzelm
parents: 11505
diff changeset
  1373
52232
a2d4ae3e04a2 tuned signature;
wenzelm
parents: 52091
diff changeset
  1374
val rewrite_rule = Conv.fconv_rule o rewrite true;
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1375
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1376
(*simple term rewriting -- no proof*)
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16380
diff changeset
  1377
fun rewrite_term thy rules procs =
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17046
diff changeset
  1378
  Pattern.rewrite_term thy (map decomp_simp' rules) procs;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1379
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1380
fun rewrite_thm mode prover ctxt = Conv.fconv_rule (rewrite_cterm mode prover ctxt);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1381
23536
60a1672e298e moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
wenzelm
parents: 23221
diff changeset
  1382
(*Rewrite the subgoals of a proof state (represented by a theorem)*)
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1383
fun rewrite_goals_rule thms th =
23584
9b5ba76de1c2 tuned goal conversion interfaces;
wenzelm
parents: 23536
diff changeset
  1384
  Conv.fconv_rule (Conv.prems_conv ~1 (rewrite_cterm (true, true, true) simple_prover
35232
f588e1169c8b renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
wenzelm
parents: 35231
diff changeset
  1385
    (global_context (Thm.theory_of_thm th) empty_ss addsimps thms))) th;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1386
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1387
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1388
(** meta-rewriting tactics **)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1389
28839
32d498cf7595 eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
wenzelm
parents: 28620
diff changeset
  1390
(*Rewrite all subgoals*)
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1391
fun rewrite_goals_tac defs = PRIMITIVE (rewrite_goals_rule defs);
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1392
28839
32d498cf7595 eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
wenzelm
parents: 28620
diff changeset
  1393
(*Rewrite one subgoal*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1394
fun generic_rewrite_goal_tac mode prover_tac ctxt i thm =
25203
e5b2dd8db7c8 asm_rewrite_goal_tac: avoiding PRIMITIVE lets informative exceptions (from simprocs) get through;
wenzelm
parents: 24707
diff changeset
  1395
  if 0 < i andalso i <= Thm.nprems_of thm then
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1396
    Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ctxt) i thm)
25203
e5b2dd8db7c8 asm_rewrite_goal_tac: avoiding PRIMITIVE lets informative exceptions (from simprocs) get through;
wenzelm
parents: 24707
diff changeset
  1397
  else Seq.empty;
23536
60a1672e298e moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
wenzelm
parents: 23221
diff changeset
  1398
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1399
fun rewrite_goal_tac rews i st =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1400
  generic_rewrite_goal_tac (true, false, false) (K no_tac)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1401
    (global_context (Thm.theory_of_thm st) empty_ss addsimps rews) i st;
23536
60a1672e298e moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
wenzelm
parents: 23221
diff changeset
  1402
46707
1427dcc7c9a6 eliminated odd comment from distant past;
wenzelm
parents: 46465
diff changeset
  1403
(*Prunes all redundant parameters from the proof state by rewriting.*)
46186
9ae331a1d8c5 more qualified names;
wenzelm
parents: 45625
diff changeset
  1404
val prune_params_tac = rewrite_goals_tac [Drule.triv_forall_equality];
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1405
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1406
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1407
(* for folding definitions, handling critical pairs *)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1408
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1409
(*The depth of nesting in a term*)
32797
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
  1410
fun term_depth (Abs (_, _, t)) = 1 + term_depth t
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
  1411
  | term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t)
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1412
  | term_depth _ = 0;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1413
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1414
val lhs_of_thm = #1 o Logic.dest_equals o prop_of;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1415
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1416
(*folding should handle critical pairs!  E.g. K == Inl(0),  S == Inr(Inl(0))
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1417
  Returns longest lhs first to avoid folding its subexpressions.*)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1418
fun sort_lhs_depths defs =
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1419
  let val keylist = AList.make (term_depth o lhs_of_thm) defs
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1420
      val keys = sort_distinct (rev_order o int_ord) (map #2 keylist)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1421
  in map (AList.find (op =) keylist) keys end;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1422
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  1423
val rev_defs = sort_lhs_depths o map Thm.symmetric;
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1424
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1425
fun fold_rule defs = fold rewrite_rule (rev_defs defs);
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1426
fun fold_goals_tac defs = EVERY (map rewrite_goals_tac (rev_defs defs));
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1427
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1428
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1429
(* HHF normal form: !! before ==>, outermost !! generalized *)
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1430
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1431
local
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1432
21565
bd28361f4c5b simplified '?' operator;
wenzelm
parents: 21516
diff changeset
  1433
fun gen_norm_hhf ss th =
bd28361f4c5b simplified '?' operator;
wenzelm
parents: 21516
diff changeset
  1434
  (if Drule.is_norm_hhf (Thm.prop_of th) then th
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1435
   else
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1436
    Conv.fconv_rule
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1437
      (rewrite_cterm (true, false, false) (K (K NONE))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1438
        (global_context (Thm.theory_of_thm th) ss)) th)
21565
bd28361f4c5b simplified '?' operator;
wenzelm
parents: 21516
diff changeset
  1439
  |> Thm.adjust_maxidx_thm ~1
bd28361f4c5b simplified '?' operator;
wenzelm
parents: 21516
diff changeset
  1440
  |> Drule.gen_all;
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1441
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1442
val hhf_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1443
  simpset_of (empty_simpset (Context.proof_of (Context.the_thread_data ()))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1444
    addsimps Drule.norm_hhf_eqs);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1445
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1446
val hhf_protect_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1447
  simpset_of (empty_simpset (Context.proof_of (Context.the_thread_data ()))
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1448
    addsimps Drule.norm_hhf_eqs |> add_eqcong Drule.protect_cong);
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1449
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1450
in
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1451
26424
a6cad32a27b0 eliminated theory ProtoPure;
wenzelm
parents: 25472
diff changeset
  1452
val norm_hhf = gen_norm_hhf hhf_ss;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1453
val norm_hhf_protect = gen_norm_hhf hhf_protect_ss;
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1454
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1455
end;
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1456
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1457
end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1458
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
  1459
structure Basic_Meta_Simplifier: BASIC_RAW_SIMPLIFIER = Raw_Simplifier;
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31298
diff changeset
  1460
open Basic_Meta_Simplifier;