src/Pure/raw_simplifier.ML
author wenzelm
Mon, 02 Dec 2024 22:16:29 +0100
changeset 81541 5335b1ca6233
parent 81534 c32ebdcbe8ca
permissions -rw-r--r--
more elementary operation Term.variant_bounds: only for bounds vs. frees, no consts, no tfrees;
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
80712
05b16602a683 clarified signature;
wenzelm
parents: 80711
diff changeset
    20
  val mk_rrules: Proof.context -> thm -> rrule list
16807
730cace0ae48 export eq_rrule;
wenzelm
parents: 16665
diff changeset
    21
  val eq_rrule: rrule * rrule -> bool
78800
0b3700d31758 clarified signature;
wenzelm
parents: 78453
diff changeset
    22
  type proc = Proof.context -> cterm -> thm option
17614
37ee526db497 added mk_solver';
wenzelm
parents: 17496
diff changeset
    23
  type solver
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    24
  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
    25
  type simpset
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    26
  val empty_ss: simpset
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    27
  val merge_ss: simpset * simpset -> simpset
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
    28
  datatype proc_kind = Simproc | Congproc of bool
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
    29
  type simproc
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
    30
  val cert_simproc: theory ->
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
    31
    {name: string, kind: proc_kind, lhss: term list,
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
    32
      proc: proc Morphism.entity, identifier: thm list} -> simproc
45290
f599ac41e7f5 tuned signature -- refined terminology;
wenzelm
parents: 44058
diff changeset
    33
  val transform_simproc: morphism -> simproc -> simproc
78114
43154a48da69 clarified treatment of context;
wenzelm
parents: 78084
diff changeset
    34
  val trim_context_simproc: simproc -> simproc
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    35
  val simpset_of: Proof.context -> simpset
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    36
  val put_simpset: simpset -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    37
  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
    38
  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
    39
  val empty_simpset: Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    40
  val clear_simpset: Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    41
  val addsimps: Proof.context * thm list -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    42
  val delsimps: Proof.context * thm list -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    43
  val addsimprocs: Proof.context * simproc list -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    44
  val delsimprocs: Proof.context * simproc list -> Proof.context
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
    45
  val setloop: Proof.context * (Proof.context -> int -> tactic) -> Proof.context
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
    46
  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
    47
  val delloop: Proof.context * string -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    48
  val setSSolver: Proof.context * solver -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    49
  val addSSolver: Proof.context * solver -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    50
  val setSolver: Proof.context * solver -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    51
  val addSolver: Proof.context * solver -> Proof.context
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
    52
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    53
  val rewrite_rule: Proof.context -> thm list -> thm -> thm
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    54
  val rewrite_goals_rule: Proof.context -> thm list -> thm -> thm
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    55
  val rewrite_goals_tac: Proof.context -> thm list -> tactic
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    56
  val rewrite_goal_tac: Proof.context -> thm list -> int -> tactic
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    57
  val prune_params_tac: Proof.context -> tactic
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    58
  val fold_rule: Proof.context -> thm list -> thm -> thm
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
    59
  val fold_goals_tac: Proof.context -> thm list -> tactic
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
    60
  val norm_hhf: Proof.context -> thm -> thm
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 54742
diff changeset
    61
  val norm_hhf_protect: Proof.context -> thm -> thm
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
    62
end;
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
    63
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
    64
signature RAW_SIMPLIFIER =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
    65
sig
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
    66
  include BASIC_RAW_SIMPLIFIER
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
    67
  exception SIMPLIFIER of string * thm list
80706
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    68
  val dest_simps: simpset -> (Thm_Name.T * thm) list
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    69
  val dest_congs: simpset -> (cong_name * thm) list
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    70
  val dest_ss: simpset ->
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    71
   {simps: (Thm_Name.T * thm) list,
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    72
    simprocs: (string * term list) list,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
    73
    congprocs: (string * {lhss: term list, proc: proc Morphism.entity}) list,
80706
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    74
    congs: (cong_name * thm) list,
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    75
    weak_congs: cong_name list,
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    76
    loopers: string list,
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    77
    unsafe_solvers: string list,
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
    78
    safe_solvers: string list}
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
    79
  val add_proc: simproc -> Proof.context -> Proof.context
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
    80
  val del_proc: simproc -> Proof.context -> Proof.context
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
    81
  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
    82
  val set_trace_ops: trace_ops -> theory -> theory
71235
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
    83
  val subgoal_tac: Proof.context -> int -> tactic
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
    84
  val loop_tac: Proof.context -> int -> tactic
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
    85
  val solvers: Proof.context -> solver list * solver list
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    86
  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
    87
  val prems_of: Proof.context -> thm list
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    88
  val add_simp: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    89
  val del_simp: thm -> Proof.context -> Proof.context
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
    90
  val flip_simp: thm -> Proof.context -> Proof.context
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
    91
  val init_simpset: thm list -> Proof.context -> Proof.context
80711
043e5fd3ce32 more direct access to Simplifier.mk_cong, to avoid odd Simpdata.mk_meta_cong seen in the wild;
wenzelm
parents: 80710
diff changeset
    92
  val mk_cong: Proof.context -> thm -> thm
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    93
  val add_eqcong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    94
  val del_eqcong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    95
  val add_cong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    96
  val del_cong: thm -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
    97
  val mksimps: Proof.context -> thm -> thm list
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
    98
  val get_mksimps_context: Proof.context -> (thm -> Proof.context -> thm list * Proof.context)
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
    99
  val set_mksimps_context: (thm -> Proof.context -> thm list * Proof.context) ->
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   100
    Proof.context -> Proof.context
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   101
  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
   102
  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
   103
  val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context
70586
57df8a85317a clarified signature;
wenzelm
parents: 70528
diff changeset
   104
  val set_term_ord: term ord -> Proof.context -> Proof.context
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   105
  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
   106
  val solver: Proof.context -> solver -> int -> tactic
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 add_prems: thm list -> Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   109
  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
   110
    Proof.context -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   111
  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
   112
  val rewrite_cterm: bool * bool * bool ->
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   113
    (Proof.context -> thm -> thm option) -> Proof.context -> conv
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16380
diff changeset
   114
  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
   115
  val rewrite_thm: bool * bool * bool ->
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   116
    (Proof.context -> thm -> thm option) -> Proof.context -> thm -> thm
46465
5ba52c337cd0 tuned signature;
wenzelm
parents: 46462
diff changeset
   117
  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
   118
    (Proof.context -> tactic) -> Proof.context -> int -> tactic
78453
3fdf3c5cfa9d performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents: 78136
diff changeset
   119
  val rewrite0: Proof.context -> bool -> conv
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
   120
  val rewrite: Proof.context -> bool -> thm list -> conv
78453
3fdf3c5cfa9d performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents: 78136
diff changeset
   121
  val rewrite0_rule: Proof.context -> thm -> thm
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   122
end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   123
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
   124
structure Raw_Simplifier: RAW_SIMPLIFIER =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   125
struct
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   126
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   127
(** datatype simpset **)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   128
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   129
(* congruence rules *)
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   130
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   131
type cong_name = bool * string;
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   132
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   133
fun cong_name (Const (a, _)) = SOME (true, a)
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   134
  | cong_name (Free (a, _)) = SOME (false, a)
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   135
  | cong_name _ = NONE;
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   136
71239
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   137
structure Congtab = Table(type key = cong_name val ord = prod_ord bool_ord fast_string_ord);
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   138
51590
c52891242de2 more formal cong_name;
wenzelm
parents: 49660
diff changeset
   139
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   140
(* rewrite rules *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   141
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   142
type rrule =
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   143
 {thm: thm,         (*the rewrite rule*)
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   144
  name: Thm_Name.T, (*name of theorem from which rewrite rule was extracted*)
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   145
  lhs: term,        (*the left-hand side*)
58836
4037bb00d08e tuned spelling;
wenzelm
parents: 57859
diff changeset
   146
  elhs: cterm,      (*the eta-contracted lhs*)
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   147
  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
   148
  fo: bool,         (*use first-order matching*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   149
  perm: bool};      (*the rewrite rule is permutative*)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   150
61057
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
   151
fun trim_context_rrule ({thm, name, lhs, elhs, extra, fo, perm}: rrule) =
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
   152
  {thm = Thm.trim_context thm, name = name, lhs = lhs, elhs = Thm.trim_context_cterm elhs,
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
   153
    extra = extra, fo = fo, perm = perm};
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
   154
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   155
(*
12603
7d2bca103101 tuned tracing markup;
wenzelm
parents: 12285
diff changeset
   156
Remarks:
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   157
  - elhs is used for matching,
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   158
    lhs only for preservation of bound variable names;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   159
  - fo is set iff
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   160
    either elhs is first-order (no Var is applied),
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   161
      in which case fo-matching is complete,
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   162
    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
   163
      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
   164
*)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   165
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   166
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
   167
  Thm.eq_thm_prop (thm1, thm2);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   168
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   169
(* 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
   170
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
   171
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
   172
*)
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   173
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
   174
  let
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   175
    val elhss = elhs :: prems;
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   176
    val tvars = TVars.build (fold TVars.add_tvars elhss);
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   177
    val vars = Vars.build (fold Vars.add_vars elhss);
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   178
  in
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   179
    erhs |> Term.exists_type (Term.exists_subtype
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   180
      (fn TVar v => not (TVars.defined tvars v) | _ => false)) orelse
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   181
    erhs |> Term.exists_subterm
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   182
      (fn Var v => not (Vars.defined vars v) | _ => false)
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   183
  end;
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   184
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   185
fun rrule_extra_vars elhs thm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   186
  rewrite_rule_extra_vars [] (Thm.term_of elhs) (Thm.full_prop_of thm);
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   187
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   188
fun mk_rrule2 {thm, name, lhs, elhs, perm} =
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   189
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   190
    val t = Thm.term_of elhs;
20546
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
   191
    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
   192
    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
   193
  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
   194
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   195
(*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
   196
fun default_reorient ctxt prems lhs rhs =
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   197
  rewrite_rule_extra_vars prems lhs rhs
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   198
    orelse
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   199
  is_Var (head_of lhs)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   200
    orelse
16305
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   201
(* 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
   202
   usually it is very useful :-(
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   203
  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
   204
  andalso not(exists_subterm is_Var lhs)
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   205
    orelse
5e7b6731b004 Added the t = x "fix" - in (* ... *) :-(
nipkow
parents: 16042
diff changeset
   206
*)
16842
wenzelm
parents: 16807
diff changeset
   207
  exists (fn t => Logic.occs (lhs, t)) (rhs :: prems)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   208
    orelse
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   209
  null prems andalso Pattern.matches (Proof_Context.theory_of ctxt) (lhs, rhs)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   210
    (*the condition "null prems" is necessary because conditional rewrites
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   211
      with extra variables in the conditions may terminate although
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
   212
      the rhs is an instance of the lhs; example: ?m < ?n \<Longrightarrow> f ?n \<equiv> f ?m *)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   213
    orelse
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   214
  is_Const lhs andalso not (is_Const rhs);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   215
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   216
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   217
(* simplification procedures *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   218
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   219
datatype proc_kind = Simproc | Congproc of bool;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   220
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   221
val is_congproc = fn Congproc _ => true | _ => false;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   222
val is_weak_congproc = fn Congproc weak => weak | _ => false;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   223
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   224
fun map_procs kind f (simprocs, congprocs) =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   225
  if is_congproc kind then (simprocs, f congprocs) else (f simprocs, congprocs);
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   226
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   227
fun print_proc_kind Simproc = "simplification procedure"
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   228
  | print_proc_kind (Congproc false) = "simplification procedure (cong)"
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   229
  | print_proc_kind (Congproc true) = "simplification procedure (weak cong)";
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   230
78800
0b3700d31758 clarified signature;
wenzelm
parents: 78453
diff changeset
   231
type proc = Proof.context -> cterm -> thm option;
0b3700d31758 clarified signature;
wenzelm
parents: 78453
diff changeset
   232
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   233
datatype 'lhs procedure =
80697
48eaf5c85d6e tuned signature: support more general procedures;
wenzelm
parents: 80306
diff changeset
   234
  Procedure of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   235
   {name: string,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   236
    kind: proc_kind,
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   237
    lhs: 'lhs,
78800
0b3700d31758 clarified signature;
wenzelm
parents: 78453
diff changeset
   238
    proc: proc Morphism.entity,
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   239
    id: stamp * thm list};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   240
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   241
fun procedure_kind (Procedure {kind, ...}) = kind;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   242
fun procedure_lhs (Procedure {lhs, ...}) = lhs;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   243
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   244
fun eq_procedure_id (Procedure {id = (s1, ths1), ...}, Procedure {id = (s2, ths2), ...}) =
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   245
  s1 = s2 andalso eq_list Thm.eq_thm_prop (ths1, ths2);
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   246
51717
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
(* solvers *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   249
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   250
datatype solver =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   251
  Solver of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   252
   {name: string,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   253
    solver: Proof.context -> int -> tactic,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   254
    id: stamp};
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   255
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   256
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
   257
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   258
fun solver_name (Solver {name, ...}) = name;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   259
fun solver ctxt (Solver {solver = tac, ...}) = tac ctxt;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   260
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
   261
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   262
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   263
(* simplification sets *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   264
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   265
(*A simpset contains data required during conversion:
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   266
    rules: discrimination net of rewrite rules;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   267
    prems: current premises;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   268
    depth: simp_depth and exceeded flag;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   269
    congs: association list of congruence rules and
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   270
           a list of `weak' congruence constants.
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   271
           A congruence is `weak' if it avoids normalization of some argument.
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   272
    procs: simplification procedures indexed via discrimination net
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   273
      simprocs: functions that prove rewrite rules on the fly;
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   274
      congprocs: functions that prove congruence rules on the fly;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   275
    mk_rews:
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   276
      mk: turn simplification thms into rewrite rules (and update context);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   277
      mk_cong: prepare congruence rules;
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
   278
      mk_sym: turn \<equiv> around;
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
   279
      mk_eq_True: turn P into P \<equiv> True;
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   280
    term_ord: for ordered rewriting;*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   281
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   282
datatype simpset =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   283
  Simpset of
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   284
   {rules: rrule Net.net,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   285
    prems: thm list,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   286
    depth: int * bool Unsynchronized.ref} *
71239
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   287
   {congs: thm Congtab.table * cong_name list,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   288
    procs: term procedure Net.net * term procedure Net.net,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   289
    mk_rews:
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   290
     {mk: thm -> Proof.context -> thm list * Proof.context,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   291
      mk_cong: Proof.context -> thm -> thm,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   292
      mk_sym: Proof.context -> thm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   293
      mk_eq_True: Proof.context -> thm -> thm option,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   294
      reorient: Proof.context -> term list -> term -> term -> bool},
70586
57df8a85317a clarified signature;
wenzelm
parents: 70528
diff changeset
   295
    term_ord: term ord,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   296
    subgoal_tac: Proof.context -> int -> tactic,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   297
    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
   298
    solvers: solver list * solver list};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   299
54728
445e7947c6b5 tuned signature;
wenzelm
parents: 54727
diff changeset
   300
fun internal_ss (Simpset (_, ss2)) = ss2;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   301
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   302
fun make_ss1 (rules, prems, depth) = {rules = rules, prems = prems, depth = depth};
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   303
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   304
fun map_ss1 f {rules, prems, depth} = make_ss1 (f (rules, prems, depth));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   305
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   306
fun make_ss2 (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   307
  {congs = congs, procs = procs, mk_rews = mk_rews, term_ord = term_ord,
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
   308
    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
   309
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   310
fun map_ss2 f {congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers} =
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   311
  make_ss2 (f (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   312
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   313
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
   314
80706
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   315
fun dest_simps (Simpset ({rules, ...}, _)) =
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   316
  Net.entries rules
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   317
  |> map (fn {name, thm, ...} => (name, thm));
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   318
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   319
fun dest_congs (Simpset (_, {congs, ...})) = Congtab.dest (#1 congs);
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   320
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   321
fun dest_procs procs =
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   322
  Net.entries procs
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   323
  |> partition_eq eq_procedure_id
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   324
  |> map (fn ps as Procedure {name, proc, ...} :: _ =>
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   325
      (name, {lhss = map (fn Procedure {lhs, ...} => lhs) ps, proc = proc}));
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   326
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   327
fun dest_ss (ss as Simpset (_, {congs, procs = (simprocs, congprocs), loop_tacs, solvers, ...})) =
80706
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   328
 {simps = dest_simps ss,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   329
  simprocs = map (apsnd #lhss) (dest_procs simprocs),
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   330
  congprocs = dest_procs congprocs,
80706
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   331
  congs = dest_congs ss,
29734511c661 clarified signature;
wenzelm
parents: 80705
diff changeset
   332
  weak_congs = #2 congs,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   333
  loopers = map fst loop_tacs,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   334
  unsafe_solvers = map solver_name (#1 solvers),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   335
  safe_solvers = map solver_name (#2 solvers)};
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
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   338
(* empty *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   339
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   340
fun init_ss depth mk_rews term_ord subgoal_tac solvers =
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   341
  make_simpset ((Net.empty, [], depth),
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   342
    ((Congtab.empty, []), (Net.empty, Net.empty), mk_rews, term_ord, subgoal_tac, [], solvers));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   343
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   344
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
   345
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   346
val empty_ss =
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   347
  init_ss (0, Unsynchronized.ref false)
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   348
    {mk = fn th => pair (if can Logic.dest_equals (Thm.concl_of th) then [th] else []),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   349
      mk_cong = K I,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   350
      mk_sym = default_mk_sym,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   351
      mk_eq_True = K (K NONE),
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   352
      reorient = default_reorient}
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   353
    Term_Ord.term_ord (K (K no_tac)) ([], []);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   354
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   355
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   356
(* merge *)  (*NOTE: ignores some fields of 2nd simpset*)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   357
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   358
fun merge_ss (ss1, ss2) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   359
  if pointer_eq (ss1, ss2) then ss1
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   360
  else
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   361
    let
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   362
      val Simpset ({rules = rules1, prems = prems1, depth = depth1},
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   363
       {congs = (congs1, weak1), procs = (simprocs1, congprocs1), mk_rews, term_ord, 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
   364
        loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1;
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   365
      val Simpset ({rules = rules2, prems = prems2, depth = depth2},
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   366
       {congs = (congs2, weak2), procs = (simprocs2, congprocs2),
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   367
        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
   368
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   369
      val rules' = Net.merge eq_rrule (rules1, rules2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   370
      val prems' = Thm.merge_thms (prems1, prems2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   371
      val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1;
71239
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   372
      val congs' = Congtab.merge (K true) (congs1, congs2);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   373
      val weak' = merge (op =) (weak1, weak2);
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   374
      val simprocs' = Net.merge eq_procedure_id (simprocs1, simprocs2);
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   375
      val congprocs' = Net.merge eq_procedure_id (congprocs1, congprocs2);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   376
      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
   377
      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
   378
      val solvers' = merge eq_solver (solvers1, solvers2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   379
    in
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   380
      make_simpset ((rules', prems', depth'), ((congs', weak'), (simprocs', congprocs'),
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   381
        mk_rews, term_ord, subgoal_tac, loop_tacs', (unsafe_solvers', solvers')))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   382
    end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   383
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   384
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   385
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   386
(** context data **)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   387
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   388
structure Simpset = Generic_Data
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   389
(
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   390
  type T = simpset;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   391
  val empty = empty_ss;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   392
  val merge = merge_ss;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   393
);
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
val simpset_of = Simpset.get o Context.Proof;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   396
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   397
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
   398
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
   399
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
   400
71403
haftmann
parents: 71318
diff changeset
   401
fun put_simpset ss = map_simpset (K ss);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   402
71403
haftmann
parents: 71318
diff changeset
   403
fun simpset_map ctxt f ss = ctxt |> put_simpset ss |> f |> simpset_of;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   404
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   405
val empty_simpset = put_simpset empty_ss;
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
fun map_theory_simpset f thy =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   408
  let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   409
    val ctxt' = f (Proof_Context.init_global thy);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   410
    val thy' = Proof_Context.theory_of ctxt';
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   411
  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
   412
57859
29e728588163 more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
wenzelm
parents: 56438
diff changeset
   413
fun map_ss f = Context.mapping (map_theory_simpset (f o Context_Position.not_really)) f;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   414
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   415
val clear_simpset =
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   416
  map_simpset (fn Simpset ({depth, ...}, {mk_rews, term_ord, subgoal_tac, solvers, ...}) =>
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   417
    init_ss depth mk_rews term_ord subgoal_tac solvers);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   418
80710
wenzelm
parents: 80709
diff changeset
   419
val get_mk_rews = simpset_of #> (fn Simpset (_, {mk_rews, ...}) => mk_rews);
wenzelm
parents: 80709
diff changeset
   420
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   421
71235
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   422
(* accessors for tactis *)
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   423
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   424
fun subgoal_tac ctxt = (#subgoal_tac o internal_ss o simpset_of) ctxt ctxt;
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   425
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   426
fun loop_tac ctxt =
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   427
  FIRST' (map (fn (_, tac) => tac ctxt) (rev ((#loop_tacs o internal_ss o simpset_of) ctxt)));
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   428
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   429
val solvers = #solvers o internal_ss o simpset_of
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   430
d12c58e12c51 more direct accessors for simpset
haftmann
parents: 71234
diff changeset
   431
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   432
(* simp depth *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   433
66934
b86513bcf7ac reduced simp_depth_limit
nipkow
parents: 64556
diff changeset
   434
(*
b86513bcf7ac reduced simp_depth_limit
nipkow
parents: 64556
diff changeset
   435
The simp_depth_limit is meant to abort infinite recursion of the simplifier
b86513bcf7ac reduced simp_depth_limit
nipkow
parents: 64556
diff changeset
   436
early but should not terminate "normal" executions.
b86513bcf7ac reduced simp_depth_limit
nipkow
parents: 64556
diff changeset
   437
As of 2017, 25 would suffice; 40 builds in a safety margin.
b86513bcf7ac reduced simp_depth_limit
nipkow
parents: 64556
diff changeset
   438
*)
b86513bcf7ac reduced simp_depth_limit
nipkow
parents: 64556
diff changeset
   439
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69137
diff changeset
   440
val simp_depth_limit = Config.declare_int ("simp_depth_limit", \<^here>) (K 40);
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69137
diff changeset
   441
val simp_trace_depth_limit = Config.declare_int ("simp_trace_depth_limit", \<^here>) (K 1);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   442
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   443
fun inc_simp_depth ctxt =
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   444
  ctxt |> map_simpset1 (fn (rules, prems, (depth, exceeded)) =>
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   445
    (rules, prems,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   446
      (depth + 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   447
        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
   448
        then Unsynchronized.ref false else exceeded)));
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   449
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   450
fun simp_depth ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   451
  let val Simpset ({depth = (depth, _), ...}, _) = simpset_of ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   452
  in depth end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   453
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
(* diagnostics *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   456
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   457
exception SIMPLIFIER of string * thm list;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   458
69575
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69137
diff changeset
   459
val simp_debug = Config.declare_bool ("simp_debug", \<^here>) (K false);
f77cc54f6d47 clarified signature: more types;
wenzelm
parents: 69137
diff changeset
   460
val simp_trace = Config.declare_bool ("simp_trace", \<^here>) (K false);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   461
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   462
fun cond_warning ctxt msg =
57859
29e728588163 more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
wenzelm
parents: 56438
diff changeset
   463
  if Context_Position.is_really_visible ctxt then warning (msg ()) else ();
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   464
55031
wenzelm
parents: 55028
diff changeset
   465
fun cond_tracing' ctxt flag msg =
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   466
  if Config.get ctxt flag then
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   467
    let
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   468
      val Simpset ({depth = (depth, exceeded), ...}, _) = simpset_of ctxt;
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   469
      val depth_limit = Config.get ctxt simp_trace_depth_limit;
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   470
    in
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   471
      if depth > depth_limit then
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   472
        if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true)
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   473
      else (tracing (enclose "[" "]" (string_of_int depth) ^ msg ()); exceeded := false)
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   474
    end
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   475
  else ();
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   476
55031
wenzelm
parents: 55028
diff changeset
   477
fun cond_tracing ctxt = cond_tracing' ctxt simp_trace;
wenzelm
parents: 55028
diff changeset
   478
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   479
fun print_term ctxt s t =
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   480
  s ^ "\n" ^ Syntax.string_of_term ctxt t;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   481
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   482
fun print_thm ctxt msg (name, th) =
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   483
  let
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   484
    val sffx =
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   485
      if Thm_Name.is_empty name then ""
81534
c32ebdcbe8ca proper context for extern operation: observe local options;
wenzelm
parents: 81243
diff changeset
   486
      else " " ^ quote (Global_Theory.print_thm_name ctxt name) ^ ":";
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   487
  in print_term ctxt (msg ^ sffx) (Thm.full_prop_of th) end;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   488
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   489
fun print_thm0 ctxt msg th = print_thm ctxt msg (Thm_Name.empty, th);
80304
wenzelm
parents: 80065
diff changeset
   490
51717
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
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   493
(** simpset operations **)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   494
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   495
(* prems *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   496
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   497
fun prems_of ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   498
  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
   499
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   500
fun add_prems ths =
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   501
  map_simpset1 (fn (rules, prems, depth) => (rules, ths @ prems, depth));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   502
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   503
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   504
(* maintain simp rules *)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   505
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   506
fun del_rrule loud (rrule as {thm, elhs, ...}) ctxt =
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   507
  ctxt |> map_simpset1 (fn (rules, prems, depth) =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   508
    (Net.delete_term eq_rrule (Thm.term_of elhs, rrule) rules, prems, depth))
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   509
  handle Net.DELETE =>
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   510
    (if not loud then ()
80304
wenzelm
parents: 80065
diff changeset
   511
     else cond_warning ctxt (fn () => print_thm0 ctxt "Rewrite rule not in simpset:" thm);
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   512
     ctxt);
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   513
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   514
fun insert_rrule (rrule as {thm, name, ...}) ctxt =
55031
wenzelm
parents: 55028
diff changeset
   515
 (cond_tracing ctxt (fn () => print_thm ctxt "Adding rewrite rule" (name, thm));
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   516
  ctxt |> map_simpset1 (fn (rules, prems, depth) =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   517
    let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   518
      val rrule2 as {elhs, ...} = mk_rrule2 rrule;
61057
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
   519
      val rules' = Net.insert_term eq_rrule (Thm.term_of elhs, trim_context_rrule rrule2) rules;
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
   520
    in (rules', prems, depth) end)
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   521
  handle Net.INSERT =>
80304
wenzelm
parents: 80065
diff changeset
   522
    (cond_warning ctxt (fn () => print_thm0 ctxt "Ignoring duplicate rewrite rule:" thm); ctxt));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   523
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   524
val vars_set = Vars.build o Vars.add_vars;
74227
fdcc7e8f95ea more scalable operations;
wenzelm
parents: 71403
diff changeset
   525
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   526
local
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   527
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   528
fun vperm (Var _, Var _) = true
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   529
  | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   530
  | 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
   531
  | vperm (t, u) = (t = u);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   532
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   533
fun var_perm (t, u) = vperm (t, u) andalso Vars.eq_set (apply2 vars_set (t, u));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   534
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   535
in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   536
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   537
fun decomp_simp thm =
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   538
  let
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   539
    val prop = Thm.prop_of thm;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   540
    val prems = Logic.strip_imp_prems prop;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   541
    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
   542
    val (lhs, rhs) = Thm.dest_equals concl handle TERM _ =>
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   543
      raise SIMPLIFIER ("Rewrite rule not a meta-equality", [thm]);
20579
4dc799edef89 Thm.dest_arg;
wenzelm
parents: 20546
diff changeset
   544
    val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs));
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   545
    val erhs = Envir.eta_contract (Thm.term_of rhs);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   546
    val perm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   547
      var_perm (Thm.term_of elhs, erhs) andalso
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   548
      not (Thm.term_of elhs aconv erhs) andalso
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   549
      not (is_Var (Thm.term_of elhs));
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   550
  in (prems, Thm.term_of lhs, elhs, Thm.term_of rhs, perm) end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   551
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   552
end;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   553
12783
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   554
fun decomp_simp' thm =
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   555
  let val (_, lhs, _, rhs, _) = decomp_simp thm in
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   556
    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
   557
    else (lhs, rhs)
12783
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   558
  end;
36ca5ea8092c interface to Pattern.rewrite_term;
wenzelm
parents: 12779
diff changeset
   559
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   560
fun mk_eq_True ctxt (thm, name) =
80710
wenzelm
parents: 80709
diff changeset
   561
  (case #mk_eq_True (get_mk_rews ctxt) ctxt thm of
wenzelm
parents: 80709
diff changeset
   562
    NONE => []
wenzelm
parents: 80709
diff changeset
   563
  | SOME eq_True =>
wenzelm
parents: 80709
diff changeset
   564
      let val (_, lhs, elhs, _, _) = decomp_simp eq_True;
wenzelm
parents: 80709
diff changeset
   565
      in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   566
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   567
(*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
   568
  in case there are extra vars on the rhs*)
52082
wenzelm
parents: 52037
diff changeset
   569
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
   570
  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
   571
    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
   572
      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
   573
    else [rrule]
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   574
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   575
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   576
fun mk_rrule ctxt (thm, name) =
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   577
  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
   578
    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
   579
    else
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   580
      (*weak test for loops*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   581
      if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (Thm.term_of elhs)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   582
      then mk_eq_True ctxt (thm, name)
52082
wenzelm
parents: 52037
diff changeset
   583
      else rrule_eq_True ctxt thm name lhs elhs rhs thm
70472
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
   584
  end |> map (fn {thm, name, lhs, elhs, perm} =>
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
   585
    {thm = Thm.trim_context thm, name = name, lhs = lhs,
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
   586
      elhs = Thm.trim_context_cterm elhs, perm = perm});
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   587
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   588
fun orient_rrule ctxt (thm, name) =
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   589
  let
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   590
    val (prems, lhs, elhs, rhs, perm) = decomp_simp thm;
80710
wenzelm
parents: 80709
diff changeset
   591
    val {reorient, mk_sym, ...} = get_mk_rews ctxt;
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   592
  in
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   593
    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
   594
    else if reorient ctxt prems lhs rhs then
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   595
      if reorient ctxt prems rhs lhs
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   596
      then mk_eq_True ctxt (thm, name)
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   597
      else
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   598
        (case mk_sym ctxt thm of
18208
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   599
          NONE => []
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   600
        | SOME thm' =>
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
   601
            let val (_, lhs', elhs', rhs', _) = decomp_simp thm'
52082
wenzelm
parents: 52037
diff changeset
   602
            in rrule_eq_True ctxt thm' name lhs' elhs' rhs' thm end)
wenzelm
parents: 52037
diff changeset
   603
    else rrule_eq_True ctxt thm name lhs elhs rhs thm
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   604
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   605
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   606
fun extract_rews sym thm ctxt =
68046
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67721
diff changeset
   607
  let
80710
wenzelm
parents: 80709
diff changeset
   608
    val mk = #mk (get_mk_rews ctxt);
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   609
    val (rews, ctxt') = mk thm ctxt;
80712
05b16602a683 clarified signature;
wenzelm
parents: 80711
diff changeset
   610
    val rews' = if sym then rews RL [Drule.symmetric_thm] else rews;
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   611
  in (map (rpair (Thm.get_name_hint thm)) rews', ctxt') end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   612
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   613
fun extract_safe_rrules thm ctxt =
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   614
  extract_rews false thm ctxt |>> maps (orient_rrule ctxt);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   615
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   616
fun mk_rrules ctxt thm =
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 55032
diff changeset
   617
  let
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   618
    val rews = #1 (extract_rews false thm ctxt);
80712
05b16602a683 clarified signature;
wenzelm
parents: 80711
diff changeset
   619
    val raw_rrules = maps (mk_rrule ctxt) rews;
80710
wenzelm
parents: 80709
diff changeset
   620
  in map mk_rrule2 raw_rrules end;
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 55032
diff changeset
   621
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   622
20028
b9752164ad92 add/del_simps: warning for inactive simpset (no context);
wenzelm
parents: 19798
diff changeset
   623
(* add/del rules explicitly *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   624
61090
16f7f9aa4263 more thorough transfer;
wenzelm
parents: 61057
diff changeset
   625
local
16f7f9aa4263 more thorough transfer;
wenzelm
parents: 61057
diff changeset
   626
68046
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67721
diff changeset
   627
fun comb_simps ctxt comb mk_rrule sym thms =
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   628
  let val rews = maps (fn thm => #1 (extract_rews sym (Thm.transfer' ctxt thm) ctxt)) thms;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   629
  in fold (fold comb o mk_rrule) rews ctxt end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   630
68405
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   631
(*
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   632
This code checks if the symetric version of a rule is already in the simpset.
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   633
However, the variable names in the two versions of the rule may differ.
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   634
Thus the current test modulo eq_rrule is too weak to be useful
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   635
and needs to be refined.
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   636
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   637
fun present ctxt rules (rrule as {thm, elhs, ...}) =
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   638
  (Net.insert_term eq_rrule (Thm.term_of elhs, trim_context_rrule rrule) rules;
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   639
   false)
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   640
  handle Net.INSERT =>
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   641
    (cond_warning ctxt
80304
wenzelm
parents: 80065
diff changeset
   642
       (fn () => print_thm0 ctxt "Symmetric rewrite rule already in simpset:" thm);
68405
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   643
     true);
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   644
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   645
fun sym_present ctxt thms =
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   646
  let
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   647
    val rews = extract_rews ctxt true (map (Thm.transfer' ctxt) thms);
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   648
    val rrules = map mk_rrule2 (flat(map (mk_rrule ctxt) rews))
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   649
    val Simpset({rules, ...},_) = simpset_of ctxt
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   650
  in exists (present ctxt rules) rrules end
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   651
*)
61090
16f7f9aa4263 more thorough transfer;
wenzelm
parents: 61057
diff changeset
   652
in
16f7f9aa4263 more thorough transfer;
wenzelm
parents: 61057
diff changeset
   653
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   654
fun ctxt addsimps thms =
68046
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67721
diff changeset
   655
  comb_simps ctxt insert_rrule (mk_rrule ctxt) false thms;
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67721
diff changeset
   656
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67721
diff changeset
   657
fun addsymsimps ctxt thms =
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67721
diff changeset
   658
  comb_simps ctxt insert_rrule (mk_rrule ctxt) true thms;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   659
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   660
fun ctxt delsimps thms =
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   661
  comb_simps ctxt (del_rrule true) (map mk_rrule2 o mk_rrule ctxt) false thms;
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   662
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   663
fun delsimps_quiet ctxt thms =
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   664
  comb_simps ctxt (del_rrule false) (map mk_rrule2 o mk_rrule ctxt) false thms;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   665
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   666
fun add_simp thm ctxt = ctxt addsimps [thm];
68405
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   667
(*
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   668
with check for presence of symmetric version:
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   669
  if sym_present ctxt [thm]
80304
wenzelm
parents: 80065
diff changeset
   670
  then (cond_warning ctxt (fn () => print_thm0 ctxt "Ignoring rewrite rule:" thm); ctxt)
68405
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   671
  else ctxt addsimps [thm];
6a0852b8e5a8 comments
nipkow
parents: 68403
diff changeset
   672
*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   673
fun del_simp thm ctxt = ctxt delsimps [thm];
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68046
diff changeset
   674
fun flip_simp thm ctxt = addsymsimps (delsimps_quiet ctxt [thm]) [thm];
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   675
61090
16f7f9aa4263 more thorough transfer;
wenzelm
parents: 61057
diff changeset
   676
end;
16f7f9aa4263 more thorough transfer;
wenzelm
parents: 61057
diff changeset
   677
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
   678
fun init_simpset thms ctxt = ctxt
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
   679
  |> Context_Position.set_visible false
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
   680
  |> empty_simpset
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
   681
  |> fold add_simp thms
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
   682
  |> Context_Position.restore_visible ctxt;
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
   683
57859
29e728588163 more careful treatment of context visibility for rule declarations (see also 39d9c7f175e0, e639d91d9073) -- avoid duplicate warnings;
wenzelm
parents: 56438
diff changeset
   684
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   685
(* congs *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   686
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   687
local
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   688
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   689
fun is_full_cong_prems [] [] = true
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   690
  | is_full_cong_prems [] _ = false
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   691
  | is_full_cong_prems (p :: prems) varpairs =
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   692
      (case Logic.strip_assums_concl p of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55635
diff changeset
   693
        Const ("Pure.eq", _) $ lhs $ rhs =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   694
          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
   695
            is_Var x andalso forall is_Bound xs andalso
20972
db0425645cc7 abandoned findrep
haftmann
parents: 20905
diff changeset
   696
            not (has_duplicates (op =) xs) andalso xs = ys andalso
20671
2aa8269a868e member (op =);
wenzelm
parents: 20579
diff changeset
   697
            member (op =) varpairs (x, y) andalso
19303
7da7e96bd74d gen_eq_set, remove (op =);
wenzelm
parents: 19142
diff changeset
   698
            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
   699
          end
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   700
      | _ => false);
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   701
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   702
fun is_full_cong thm =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   703
  let
43597
b4a093e755db tuned signature;
wenzelm
parents: 43596
diff changeset
   704
    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
   705
    val (lhs, rhs) = Logic.dest_equals concl;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   706
    val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   707
  in
20972
db0425645cc7 abandoned findrep
haftmann
parents: 20905
diff changeset
   708
    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
   709
    is_full_cong_prems prems (xs ~~ ys)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   710
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   711
80711
043e5fd3ce32 more direct access to Simplifier.mk_cong, to avoid odd Simpdata.mk_meta_cong seen in the wild;
wenzelm
parents: 80710
diff changeset
   712
in
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45405
diff changeset
   713
80711
043e5fd3ce32 more direct access to Simplifier.mk_cong, to avoid odd Simpdata.mk_meta_cong seen in the wild;
wenzelm
parents: 80710
diff changeset
   714
fun mk_cong ctxt = #mk_cong (get_mk_rews ctxt) ctxt;
45620
f2a587696afb modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45405
diff changeset
   715
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   716
fun add_eqcong thm ctxt = ctxt |> map_simpset2
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   717
  (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   718
    let
45621
wenzelm
parents: 45620
diff changeset
   719
      val (lhs, _) = Logic.dest_equals (Thm.concl_of thm)
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   720
        handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]);
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18573
diff changeset
   721
    (*val lhs = Envir.eta_contract lhs;*)
45621
wenzelm
parents: 45620
diff changeset
   722
      val a = the (cong_name (head_of lhs)) handle Option.Option =>
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   723
        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
   724
      val (xs, weak) = congs;
71239
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   725
      val xs' = Congtab.update (a, Thm.trim_context thm) xs;
22221
8a8aa6114a89 changed cong alist - now using AList operations instead of overwrite_warn
haftmann
parents: 22008
diff changeset
   726
      val weak' = if is_full_cong thm then weak else a :: weak;
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   727
    in ((xs', weak'), procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) end);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   728
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   729
fun del_eqcong thm ctxt = ctxt |> map_simpset2
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   730
  (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   731
    let
45621
wenzelm
parents: 45620
diff changeset
   732
      val (lhs, _) = Logic.dest_equals (Thm.concl_of thm)
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   733
        handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]);
18929
d81435108688 Envir.(beta_)eta_contract;
wenzelm
parents: 18573
diff changeset
   734
    (*val lhs = Envir.eta_contract lhs;*)
20057
058e913bac71 tuned exception handling;
wenzelm
parents: 20028
diff changeset
   735
      val a = the (cong_name (head_of lhs)) handle Option.Option =>
54997
811c35e81ac5 tuned signature;
wenzelm
parents: 54984
diff changeset
   736
        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
   737
      val (xs, _) = congs;
71239
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   738
      val xs' = Congtab.delete_safe a xs;
acc6cb1a1a67 proper table data structure
haftmann
parents: 71235
diff changeset
   739
      val weak' = Congtab.fold (fn (a, th) => if is_full_cong th then I else insert (op =) a) xs' [];
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   740
    in ((xs', weak'), procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) end);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   741
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   742
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
   743
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
   744
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   745
end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   746
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   747
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   748
(* simprocs *)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   749
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   750
type simproc = term list procedure;
22008
bfc462bfc574 added mk_simproc': tuned interface;
wenzelm
parents: 21962
diff changeset
   751
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   752
fun cert_simproc thy {name, kind, lhss, proc, identifier} : simproc =
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   753
  Procedure
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   754
   {name = name,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   755
    kind = kind,
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   756
    lhs = map (Sign.cert_term thy) lhss,
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   757
    proc = proc,
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   758
    id = (stamp (), map (Thm.transfer thy) identifier)};
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 61098
diff changeset
   759
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   760
fun transform_simproc phi (Procedure {name, kind, lhs, proc, id = (stamp, identifier)}) : simproc =
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   761
  Procedure
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   762
   {name = name,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   763
    kind = kind,
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   764
    lhs = map (Morphism.term phi) lhs,
78114
43154a48da69 clarified treatment of context;
wenzelm
parents: 78084
diff changeset
   765
    proc = Morphism.transform phi proc,
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   766
    id = (stamp, Morphism.fact phi identifier)};
78114
43154a48da69 clarified treatment of context;
wenzelm
parents: 78084
diff changeset
   767
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   768
fun trim_context_simproc (Procedure {name, kind, lhs, proc, id = (stamp, identifier)}) : simproc =
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   769
  Procedure
78114
43154a48da69 clarified treatment of context;
wenzelm
parents: 78084
diff changeset
   770
   {name = name,
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   771
    kind = kind,
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   772
    lhs = lhs,
78114
43154a48da69 clarified treatment of context;
wenzelm
parents: 78084
diff changeset
   773
    proc = Morphism.entity_reset_context proc,
78812
d769a183d51d simprocs may be distinguished via 'identifier': only works for ML antiquotation (see also 13252110a6fe);
wenzelm
parents: 78800
diff changeset
   774
    id = (stamp, map Thm.trim_context identifier)};
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   775
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   776
local
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   777
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   778
fun add_proc1 (proc as Procedure {name, kind, lhs, ...}) ctxt =
55031
wenzelm
parents: 55028
diff changeset
   779
 (cond_tracing ctxt (fn () =>
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   780
    print_term ctxt ("Adding " ^ print_proc_kind kind ^ " " ^ quote name ^ " for") lhs);
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   781
  ctxt |> map_simpset2
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   782
    (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   783
      (congs, map_procs kind (Net.insert_term eq_procedure_id (lhs, proc)) procs,
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   784
        mk_rews, term_ord, subgoal_tac, loop_tacs, solvers))
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   785
  handle Net.INSERT =>
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   786
    (cond_warning ctxt (fn () =>
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   787
      "Ignoring duplicate " ^ print_proc_kind kind ^ " " ^ quote name);
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   788
      ctxt));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   789
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   790
fun del_proc1 (proc as Procedure {name, kind, lhs, ...}) ctxt =
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   791
  ctxt |> map_simpset2
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   792
    (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   793
      (congs, map_procs kind (Net.delete_term eq_procedure_id (lhs, proc)) procs,
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   794
        mk_rews, term_ord, subgoal_tac, loop_tacs, solvers))
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   795
  handle Net.DELETE =>
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   796
    (cond_warning ctxt (fn () => "No " ^ print_proc_kind kind ^ " " ^ quote name ^ " in simpset");
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   797
      ctxt);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   798
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   799
fun split_proc (Procedure {name, kind, lhs = lhss, proc, id} : simproc) =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   800
  lhss |> map (fn lhs => Procedure {name = name, kind = kind, lhs = lhs, proc = proc, id = id});
22234
52ba19aaa9c2 type simproc: explicitl dependency of morphism;
wenzelm
parents: 22221
diff changeset
   801
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   802
in
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   803
80700
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   804
val add_proc = fold add_proc1 o split_proc;
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   805
val del_proc = fold del_proc1 o split_proc;
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   806
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   807
fun ctxt addsimprocs ps = fold add_proc ps ctxt;
f6c6d0988fba clarified signature: less redundant types;
wenzelm
parents: 80699
diff changeset
   808
fun ctxt delsimprocs ps = fold del_proc ps ctxt;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   809
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   810
end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   811
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   812
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   813
(* mk_rews *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   814
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   815
local
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   816
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   817
fun map_mk_rews f =
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   818
  map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   819
    let
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   820
      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
   821
      val (mk', mk_cong', mk_sym', mk_eq_True', reorient') =
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   822
        f (mk, mk_cong, mk_sym, mk_eq_True, reorient);
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   823
      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
   824
        reorient = reorient'};
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   825
    in (congs, procs, mk_rews', term_ord, subgoal_tac, loop_tacs, solvers) end);
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   826
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   827
in
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   828
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   829
val get_mksimps_context = #mk o get_mk_rews;
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   830
fun mksimps ctxt thm = #1 (get_mksimps_context ctxt thm ctxt);
30318
3d03190d2864 replaced archaic use of rep_ss by Simplifier.mksimps;
wenzelm
parents: 29269
diff changeset
   831
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
   832
fun set_mksimps_context 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
   833
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   834
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   835
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
   836
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   837
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   838
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
   839
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   840
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   841
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
   842
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   843
dbdcf366db53 simpset: added reorient field, set_reorient;
wenzelm
parents: 17966
diff changeset
   844
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
   845
  (mk, mk_cong, mk_sym, mk_eq_True, reorient));
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   846
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   847
end;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   848
14242
ec70653a02bf Added access to the mk_rews field (and friends).
skalberg
parents: 14040
diff changeset
   849
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   850
(* term_ord *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   851
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
   852
fun set_term_ord term_ord =
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   853
  map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) =>
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   854
   (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   855
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   856
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   857
(* tactics *)
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   858
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45621
diff changeset
   859
fun set_subgoaler subgoal_tac =
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   860
  map_simpset2 (fn (congs, procs, mk_rews, term_ord, _, loop_tacs, solvers) =>
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   861
   (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   862
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
   863
fun ctxt setloop tac = ctxt |>
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   864
  map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, _, solvers) =>
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   865
   (congs, procs, mk_rews, term_ord, subgoal_tac, [("", tac)], solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   866
52037
837211662fb8 tuned signature -- depend on context by default;
wenzelm
parents: 51717
diff changeset
   867
fun ctxt addloop (name, tac) = ctxt |>
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   868
  map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   869
    (congs, procs, mk_rews, term_ord, 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
   870
     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
   871
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   872
fun ctxt delloop name = ctxt |>
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   873
  map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) =>
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   874
    (congs, procs, mk_rews, term_ord, subgoal_tac,
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
   875
     (if AList.defined (op =) loop_tacs name then ()
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   876
      else cond_warning ctxt (fn () => "No such looper in simpset: " ^ quote name);
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   877
      AList.delete (op =) name loop_tacs), solvers));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   878
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   879
fun ctxt setSSolver solver = ctxt |> map_simpset2
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   880
  (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, _)) =>
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   881
    (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, [solver])));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   882
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   883
fun ctxt addSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord,
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   884
  subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, term_ord,
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
   885
    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
   886
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   887
fun ctxt setSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord,
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   888
  subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, term_ord,
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
   889
    subgoal_tac, loop_tacs, ([solver], solvers)));
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   890
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   891
fun ctxt addSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord,
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   892
  subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, term_ord,
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
   893
    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
   894
80705
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   895
fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, term_ord,
552cdee5cd43 tuned signature (again): anticipate different kinds of procs;
wenzelm
parents: 80700
diff changeset
   896
  subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, term_ord,
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
   897
  subgoal_tac, loop_tacs, (solvers, solvers)));
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   898
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   899
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   900
(* trace operations *)
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   901
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
   902
type trace_ops =
80065
60b6c735b5d5 clarified signature: prefer authentic cterm used in Simplifier, avoid potential re-certification in user-code;
wenzelm
parents: 80064
diff changeset
   903
 {trace_invoke: {depth: int, cterm: cterm} -> Proof.context -> Proof.context,
60b6c735b5d5 clarified signature: prefer authentic cterm used in Simplifier, avoid potential re-certification in user-code;
wenzelm
parents: 80064
diff changeset
   904
  trace_rrule: {unconditional: bool, cterm: cterm, thm: thm, rrule: rrule} ->
79737
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
   905
    Proof.context -> (Proof.context -> (thm * term) option) -> (thm * term) option,
79738
8ae4fc4692e8 tuned name
nipkow
parents: 79737
diff changeset
   906
  trace_simproc: {name: string, cterm: cterm} ->
79737
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
   907
    Proof.context -> (Proof.context -> thm option) -> thm option};
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
   908
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
   909
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
   910
(
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
   911
  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
   912
  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
   913
   {trace_invoke = fn _ => fn ctxt => ctxt,
80064
0d94dd2fd2d0 clarified names (see also 9c00a46d69d0, c5cd7a58cf2d);
wenzelm
parents: 79738
diff changeset
   914
    trace_rrule = fn _ => fn ctxt => fn cont => cont ctxt,
79737
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
   915
    trace_simproc = fn _ => fn ctxt => fn cont => cont 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
   916
  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
   917
);
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
   918
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
   919
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
   920
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
   921
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
   922
fun trace_invoke args ctxt = #trace_invoke (trace_ops ctxt) args ctxt;
80064
0d94dd2fd2d0 clarified names (see also 9c00a46d69d0, c5cd7a58cf2d);
wenzelm
parents: 79738
diff changeset
   923
fun trace_rrule args ctxt = #trace_rrule (trace_ops ctxt) args ctxt;
79737
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
   924
fun trace_simproc args ctxt = #trace_simproc (trace_ops ctxt) args ctxt;
15006
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   925
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   926
107e4dfd3b96 Merging the meta-simplifier with the Provers-simplifier. Next step:
skalberg
parents: 15001
diff changeset
   927
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   928
(** rewriting **)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   929
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   930
(*
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   931
  Uses conversions, see:
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   932
    L C Paulson, A higher-order implementation of rewriting,
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   933
    Science of Computer Programming 3 (1983), pages 119-149.
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   934
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   935
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   936
fun check_conv ctxt msg thm thm' =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   937
  let
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
   938
    val thm'' = Thm.transitive thm thm' handle THM _ =>
59690
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   939
      let
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   940
        val nthm' =
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   941
          Thm.transitive (Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm'
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   942
      in Thm.transitive thm nthm' handle THM _ =>
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   943
           let
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   944
             val nthm =
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   945
               Thm.transitive thm (Drule.beta_eta_conversion (Thm.rhs_of thm))
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   946
           in Thm.transitive nthm nthm' end
46b635624feb rhs of eqn is only eta- but not beta-eta-contracted; hence the latter is performed explicitly if needed
nipkow
parents: 59647
diff changeset
   947
      end
80304
wenzelm
parents: 80065
diff changeset
   948
    val _ = if msg then cond_tracing ctxt (fn () => print_thm0 ctxt "SUCCEEDED" thm') else ();
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   949
  in SOME thm'' end
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   950
  handle THM _ =>
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   951
    let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26424
diff changeset
   952
      val _ $ _ $ prop0 = Thm.prop_of thm;
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   953
      val _ =
55032
b49366215417 back to conditional tracing instead of noisy warning (see also 00e849f5b397): these incidents happen occasionally;
wenzelm
parents: 55031
diff changeset
   954
        cond_tracing ctxt (fn () =>
80304
wenzelm
parents: 80065
diff changeset
   955
          print_thm0 ctxt "Proved wrong theorem (bad subgoaler?)" thm' ^ "\n" ^
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   956
          print_term ctxt "Should have proved:" prop0);
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
   957
    in NONE end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   958
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   959
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   960
(* mk_procrule *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   961
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
   962
fun mk_procrule ctxt thm =
70528
9b3610fe74d6 treat simproc results as atomic -- more compact proof terms;
wenzelm
parents: 70472
diff changeset
   963
  let
9b3610fe74d6 treat simproc results as atomic -- more compact proof terms;
wenzelm
parents: 70472
diff changeset
   964
    val (prems, lhs, elhs, rhs, _) = decomp_simp thm
9b3610fe74d6 treat simproc results as atomic -- more compact proof terms;
wenzelm
parents: 70472
diff changeset
   965
    val thm' = Thm.close_derivation \<^here> thm;
9b3610fe74d6 treat simproc results as atomic -- more compact proof terms;
wenzelm
parents: 70472
diff changeset
   966
  in
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   967
    if rewrite_rule_extra_vars prems lhs rhs
80304
wenzelm
parents: 80065
diff changeset
   968
    then (cond_warning ctxt (fn () => print_thm0 ctxt "Extra vars on rhs:" thm); [])
80306
c2537860ccf8 more accurate thm "name_hint", using Thm_Name.T;
wenzelm
parents: 80304
diff changeset
   969
    else [mk_rrule2 {thm = thm', name = Thm_Name.empty, lhs = lhs, elhs = elhs, perm = false}]
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   970
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   971
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   972
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   973
(* rewritec: conversion to apply the meta simpset to a term *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   974
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   975
(*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
   976
  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
   977
  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
   978
  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
   979
  already in normal form.
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   980
  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
   981
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   982
val skel0 = Bound 0;
81243
wenzelm
parents: 80713
diff changeset
   983
val skel_fun = fn skel $ _ => skel | _ => skel0;
wenzelm
parents: 80713
diff changeset
   984
val skel_arg = fn _ $ skel => skel | _ => skel0;
wenzelm
parents: 80713
diff changeset
   985
val skel_body = fn Abs (_, _, skel) => skel | _ => skel0;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   986
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
   987
(*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
   988
  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
   989
  in the lhs.*)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
   990
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   991
fun weak_cong weak lhs =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   992
  if null weak then false  (*optimization*)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   993
  else exists_subterm
51591
e4aeb102ad70 amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
wenzelm
parents: 51590
diff changeset
   994
    (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
   995
      | Free (a, _) => member (op =) weak (false, a)
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   996
      | _ => false) lhs
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   997
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   998
fun uncond_skel ((_, weak), congprocs, (lhs, rhs)) =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
   999
  if weak_cong weak lhs then skel0
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1000
  else if Net.is_empty congprocs then rhs  (*optimization*)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1001
  else if exists (is_weak_congproc o procedure_kind) (Net.match_term congprocs lhs) then skel0
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1002
  else rhs;
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1003
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1004
(*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
  1005
  Otherwise those vars may become instantiated with unnormalized terms
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1006
  while the premises are solved.*)
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1007
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1008
fun cond_skel (args as (_, _, (lhs, rhs))) =
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
  1009
  if Vars.subset (vars_set rhs, vars_set lhs) then uncond_skel args
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1010
  else skel0;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1011
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1012
(*
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1013
  Rewriting -- we try in order:
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1014
    (1) beta reduction
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1015
    (2) unconditional rewrite rules
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1016
    (3) conditional rewrite rules
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1017
    (4) simplification procedures
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1018
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1019
  IMPORTANT: rewrite rules must not introduce new Vars or TVars!
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1020
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1021
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1022
fun rewritec (prover, maxt) ctxt t =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1023
  let
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1024
    val Simpset ({rules, ...}, {congs, procs = (simprocs, congprocs), term_ord, ...}) =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1025
      simpset_of ctxt;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1026
    val eta_thm = Thm.eta_conversion t;
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1027
    val eta_t' = Thm.rhs_of eta_thm;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1028
    val eta_t = Thm.term_of eta_t';
55316
885500f4aa6a interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
Lars Hupel <lars.hupel@mytum.de>
parents: 55032
diff changeset
  1029
    fun rew rrule =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1030
      let
61057
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
  1031
        val {thm = thm0, name, lhs, elhs = elhs0, extra, fo, perm} = rrule;
78136
wenzelm
parents: 78114
diff changeset
  1032
        val thm = Thm.transfer' ctxt thm0;
wenzelm
parents: 78114
diff changeset
  1033
        val elhs = Thm.transfer_cterm' ctxt elhs0;
32797
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
  1034
        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
  1035
        val (rthm, elhs') =
8923deb735ad rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
wenzelm
parents: 20330
diff changeset
  1036
          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
  1037
          else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs);
61057
5f6a1e31f3ad trim context for persistent storage;
wenzelm
parents: 60822
diff changeset
  1038
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1039
        val insts =
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1040
          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
  1041
          else Thm.match (elhs', eta_t');
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1042
        val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm);
14643
wenzelm
parents: 14330
diff changeset
  1043
        val prop' = Thm.prop_of thm';
74509
f24ade4ff3cc clarified signature;
wenzelm
parents: 74282
diff changeset
  1044
        val unconditional = Logic.no_prems prop';
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1045
        val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop');
80065
60b6c735b5d5 clarified signature: prefer authentic cterm used in Simplifier, avoid potential re-certification in user-code;
wenzelm
parents: 80064
diff changeset
  1046
        val trace_args = {unconditional = unconditional, cterm = eta_t', thm = thm', rrule = rrule};
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1047
      in
67561
f0b11413f1c9 clarified signature: prefer proper order operation;
wenzelm
parents: 66934
diff changeset
  1048
        if perm andalso is_greater_equal (term_ord (rhs', lhs'))
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1049
        then
55031
wenzelm
parents: 55028
diff changeset
  1050
         (cond_tracing ctxt (fn () =>
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1051
            print_thm ctxt "Cannot apply permutative rewrite rule" (name, thm) ^ "\n" ^
80304
wenzelm
parents: 80065
diff changeset
  1052
            print_thm0 ctxt "Term does not become smaller:" thm');
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1053
          NONE)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1054
        else
55031
wenzelm
parents: 55028
diff changeset
  1055
         (cond_tracing ctxt (fn () =>
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1056
            print_thm ctxt "Applying instance of rewrite rule" (name, thm));
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1057
          if unconditional
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1058
          then
80304
wenzelm
parents: 80065
diff changeset
  1059
           (cond_tracing ctxt (fn () => print_thm0 ctxt "Rewriting:" thm');
80064
0d94dd2fd2d0 clarified names (see also 9c00a46d69d0, c5cd7a58cf2d);
wenzelm
parents: 79738
diff changeset
  1060
            trace_rrule trace_args ctxt (fn ctxt' =>
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1061
              let
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1062
                val lr = Logic.dest_equals prop;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1063
                val SOME thm'' = check_conv ctxt' false eta_thm thm';
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1064
              in SOME (thm'', uncond_skel (congs, congprocs, lr)) end))
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1065
          else
80304
wenzelm
parents: 80065
diff changeset
  1066
           (cond_tracing ctxt (fn () => print_thm0 ctxt "Trying to rewrite:" thm');
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1067
            if simp_depth ctxt > Config.get ctxt simp_depth_limit
55031
wenzelm
parents: 55028
diff changeset
  1068
            then (cond_tracing ctxt (fn () => "simp_depth_limit exceeded - giving up"); NONE)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1069
            else
80064
0d94dd2fd2d0 clarified names (see also 9c00a46d69d0, c5cd7a58cf2d);
wenzelm
parents: 79738
diff changeset
  1070
              trace_rrule trace_args ctxt (fn ctxt' =>
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1071
                (case prover ctxt' thm' of
80304
wenzelm
parents: 80065
diff changeset
  1072
                  NONE => (cond_tracing ctxt' (fn () => print_thm0 ctxt' "FAILED" thm'); NONE)
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1073
                | SOME thm2 =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1074
                    (case check_conv ctxt' true eta_thm thm2 of
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1075
                      NONE => NONE
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1076
                    | SOME thm2' =>
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1077
                        let
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1078
                          val concl = Logic.strip_imp_concl prop;
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1079
                          val lr = Logic.dest_equals concl;
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1080
                        in SOME (thm2', cond_skel (congs, congprocs, lr)) end)))))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1081
      end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1082
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1083
    fun rews [] = NONE
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1084
      | rews (rrule :: rrules) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1085
          let val opt = rew rrule handle Pattern.MATCH => NONE
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1086
          in (case opt of NONE => rews rrules | some => some) end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1087
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1088
    fun sort_rrules rrs =
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1089
      let
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1090
        fun is_simple ({thm, ...}: rrule) =
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1091
          (case Thm.prop_of thm of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55635
diff changeset
  1092
            Const ("Pure.eq", _) $ _ $ _ => true
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1093
          | _ => false);
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1094
        fun sort [] (re1, re2) = re1 @ re2
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1095
          | sort (rr :: rrs) (re1, re2) =
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1096
              if is_simple rr
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1097
              then sort rrs (rr :: re1, re2)
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1098
              else sort rrs (re1, rr :: re2);
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1099
      in sort rrs ([], []) end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1100
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1101
    fun proc_rews [] = NONE
80697
48eaf5c85d6e tuned signature: support more general procedures;
wenzelm
parents: 80306
diff changeset
  1102
      | proc_rews (Procedure {name, proc, lhs, ...} :: ps) =
61098
e1b4b24f2ebd eliminated pointless cterms;
wenzelm
parents: 61095
diff changeset
  1103
          if Pattern.matches (Proof_Context.theory_of ctxt) (lhs, Thm.term_of t) then
55031
wenzelm
parents: 55028
diff changeset
  1104
            (cond_tracing' ctxt simp_debug (fn () =>
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1105
              print_term ctxt ("Trying procedure " ^ quote name ^ " on:") eta_t);
79737
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
  1106
             (let
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
  1107
                val ctxt' = Config.put simp_trace (Config.get ctxt simp_debug) ctxt
79738
8ae4fc4692e8 tuned name
nipkow
parents: 79737
diff changeset
  1108
                val res = trace_simproc {name = name, cterm = eta_t'} ctxt'
79737
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
  1109
                  (fn ctxt'' => Morphism.form_context' ctxt'' proc eta_t')
9c00a46d69d0 new simplifier trace_op for tracing simproc calls
nipkow
parents: 79731
diff changeset
  1110
              in case res of
79731
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1111
                NONE => (cond_tracing' ctxt simp_debug (fn () => "FAILED"); proc_rews ps)
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1112
              | SOME raw_thm =>
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1113
                  (cond_tracing ctxt (fn () =>
80304
wenzelm
parents: 80065
diff changeset
  1114
                     print_thm0 ctxt ("Procedure " ^ quote name ^ " produced rewrite rule:") raw_thm);
79731
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1115
                   (case rews (mk_procrule ctxt raw_thm) of
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1116
                     NONE =>
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1117
                      (cond_tracing ctxt (fn () =>
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1118
                         print_term ctxt ("IGNORED result of simproc " ^ quote name ^
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1119
                             " -- does not match") (Thm.term_of t));
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1120
                       proc_rews ps)
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1121
                   | some => some))
6dbe7910dcfc simplifier: no trace info from simprocs unless simp_debug = true.
nipkow
parents: 78814
diff changeset
  1122
              end))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1123
          else proc_rews ps;
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1124
  in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1125
    (case eta_t of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1126
      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
  1127
    | _ =>
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1128
      (case rews (sort_rrules (Net.match_term rules eta_t)) of
80699
34db40261287 clarified signature;
wenzelm
parents: 80698
diff changeset
  1129
        NONE => proc_rews (Net.match_term simprocs eta_t)
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1130
      | some => some))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1131
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1132
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1133
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1134
(* apply congprocs *)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1135
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1136
(* pattern order:
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1137
   p1 GREATER p2: p1 is more general than p2, p1 matches p2 but not vice versa
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1138
   p1 LESS    p2: p1 is more specific than p2, p2 matches p1 but not vice versa
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1139
   p1 EQUAL   p2: both match each other or neither match each other
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1140
*)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1141
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1142
fun pattern_order thy =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1143
  let
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1144
    fun matches arg = can (Pattern.match thy arg) (Vartab.empty, Vartab.empty);
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1145
  in
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1146
    fn (p1, p2) =>
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1147
      if matches (p1, p2) then
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1148
        if matches (p2, p1) then EQUAL
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1149
        else GREATER
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1150
      else
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1151
        if matches (p2, p1) then LESS
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1152
        else EQUAL
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1153
  end;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1154
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1155
fun app_congprocs ctxt ct =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1156
  let
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1157
    val thy = Proof_Context.theory_of ctxt;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1158
    val Simpset (_, {procs = (_, congprocs), ...}) = simpset_of ctxt;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1159
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1160
    val eta_ct = Thm.rhs_of (Thm.eta_conversion ct);
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1161
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1162
    fun proc_congs [] = NONE
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1163
      | proc_congs (Procedure {name, lhs, proc, ...} :: ps) =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1164
          if Pattern.matches thy (lhs, Thm.term_of ct) then
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1165
            let
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1166
              val _ =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1167
                cond_tracing' ctxt simp_debug (fn () =>
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1168
                  print_term ctxt ("Trying procedure " ^ quote name ^ " on:") (Thm.term_of eta_ct));
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1169
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1170
              val ctxt' = Config.put simp_trace (Config.get ctxt simp_debug) ctxt;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1171
              val res =
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1172
                trace_simproc {name = name, cterm = eta_ct} ctxt'
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1173
                  (fn ctxt'' => Morphism.form_context' ctxt'' proc eta_ct);
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1174
            in
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1175
              (case res of
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1176
                NONE => (cond_tracing' ctxt simp_debug (fn () => "FAILED"); proc_congs ps)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1177
              | SOME raw_thm =>
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1178
                  (cond_tracing ctxt (fn () =>
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1179
                     print_thm0 ctxt ("Procedure " ^ quote name ^ " produced congruence rule:")
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1180
                       raw_thm);
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1181
                   SOME (raw_thm, skel0)))
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1182
            end
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1183
          else proc_congs ps;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1184
  in
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1185
    Net.match_term congprocs (Thm.term_of eta_ct)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1186
    |> sort (pattern_order thy o apply2 procedure_lhs)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1187
    |> proc_congs
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1188
  end;
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1189
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1190
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1191
(* conversion to apply a congruence rule to a term *)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1192
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1193
fun congc prover ctxt maxt cong t =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1194
  let
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1195
    val rthm = Thm.incr_indexes (maxt + 1) cong;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1196
    val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (Thm.cprop_of rthm)));
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1197
    val insts = Thm.match (rlhs, t)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1198
    (* Thm.match can raise Pattern.MATCH;
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1199
       is handled when congc is called *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1200
    val thm' =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1201
      Thm.instantiate insts (Thm.rename_boundvars (Thm.term_of rlhs) (Thm.term_of t) rthm);
80304
wenzelm
parents: 80065
diff changeset
  1202
    val _ = cond_tracing ctxt (fn () => print_thm0 ctxt "Applying congruence rule:" thm');
wenzelm
parents: 80065
diff changeset
  1203
    fun err (msg, thm) = (cond_tracing ctxt (fn () => print_thm0 ctxt msg thm); NONE);
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1204
  in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1205
    (case prover thm' of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1206
      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
  1207
    | SOME thm2 =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1208
        (case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1209
          NONE => err ("Congruence proof failed.  Should not have proved", thm2)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1210
        | SOME thm2' =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1211
            if op aconv (apply2 Thm.term_of (Thm.dest_equals (Thm.cprop_of thm2')))
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1212
            then NONE else SOME thm2'))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1213
  end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1214
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60324
diff changeset
  1215
val vA = (("A", 0), propT);
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60324
diff changeset
  1216
val vB = (("B", 0), propT);
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60324
diff changeset
  1217
val vC = (("C", 0), propT);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1218
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1219
fun transitive1 NONE NONE = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1220
  | transitive1 (SOME thm1) NONE = SOME thm1
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1221
  | transitive1 NONE (SOME thm2) = SOME thm2
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1222
  | transitive1 (SOME thm1) (SOME thm2) = SOME (Thm.transitive thm1 thm2);
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1223
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1224
fun transitive2 thm = transitive1 (SOME thm);
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1225
fun transitive3 thm = transitive1 thm o SOME;
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1226
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1227
fun bottomc ((simprem, useprem, mutsimp), prover, maxidx) =
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1228
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1229
    fun botc skel ctxt t =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1230
      if is_Var skel then NONE
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1231
      else
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1232
        (case subc skel ctxt t of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1233
           some as SOME thm1 =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1234
             (case rewritec (prover, maxidx) ctxt (Thm.rhs_of thm1) of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1235
                SOME (thm2, skel2) =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1236
                  transitive2 (Thm.transitive thm1 thm2)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1237
                    (botc skel2 ctxt (Thm.rhs_of thm2))
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1238
              | NONE => some)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1239
         | NONE =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1240
             (case rewritec (prover, maxidx) ctxt t of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1241
                SOME (thm2, skel2) => transitive2 thm2
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1242
                  (botc skel2 ctxt (Thm.rhs_of thm2))
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1243
              | NONE => NONE))
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1244
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1245
    and try_botc ctxt t =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1246
      (case botc skel0 ctxt t of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1247
        SOME trec1 => trec1
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1248
      | NONE => Thm.reflexive t)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1249
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1250
    and subc skel ctxt t0 =
55014
a93f496f6c30 general notion of auxiliary bounds within context;
wenzelm
parents: 55000
diff changeset
  1251
        let val Simpset (_, {congs, ...}) = simpset_of ctxt in
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1252
          (case Thm.term_of t0 of
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74518
diff changeset
  1253
              Abs (a, _, _) =>
81243
wenzelm
parents: 80713
diff changeset
  1254
                let val ((v, t'), ctxt') = Variable.dest_abs_cterm t0 ctxt in
wenzelm
parents: 80713
diff changeset
  1255
                  (case botc (skel_body skel) ctxt' t' of
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 74518
diff changeset
  1256
                    SOME thm => SOME (Thm.abstract_rule a v thm)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1257
                  | NONE => NONE)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1258
                end
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1259
            | t $ _ =>
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1260
              (case t of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55635
diff changeset
  1261
                Const ("Pure.imp", _) $ _  => impc t0 ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1262
              | Abs _ =>
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1263
                  let val thm = Thm.beta_conversion false t0
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1264
                  in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1265
                    (case subc skel0 ctxt (Thm.rhs_of thm) of
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1266
                      NONE => SOME thm
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1267
                    | SOME thm' => SOME (Thm.transitive thm thm'))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1268
                  end
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1269
              | _  =>
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1270
                  let
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1271
                    fun appc () =
81243
wenzelm
parents: 80713
diff changeset
  1272
                      let val (ct, cu) = Thm.dest_comb t0 in
wenzelm
parents: 80713
diff changeset
  1273
                        (case botc (skel_fun skel) ctxt ct of
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1274
                          SOME thm1 =>
81243
wenzelm
parents: 80713
diff changeset
  1275
                            (case botc (skel_arg skel) ctxt cu of
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1276
                              SOME thm2 => SOME (Thm.combination thm1 thm2)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1277
                            | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu)))
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1278
                        | NONE =>
81243
wenzelm
parents: 80713
diff changeset
  1279
                            (case botc (skel_arg skel) ctxt cu of
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1280
                              SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1)
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1281
                            | NONE => NONE))
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1282
                      end;
80707
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1283
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1284
                    val (h, ts) = strip_comb t;
80707
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1285
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1286
     (*Prefer congprocs over plain cong rules. In congprocs prefer most specific rules.
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1287
       If there is a matching congproc, then look into the result:
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1288
         1. plain equality: consider normalisation complete (just as with a plain congruence rule),
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1289
         2. conditional rule: treat like congruence rules like SOME cong case below.*)
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1290
80707
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1291
                    fun app_cong () =
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1292
                      (case app_congprocs ctxt t0 of
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1293
                        SOME (thm, _) => SOME thm
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1294
                      | NONE => Option.mapPartial (Congtab.lookup (fst congs)) (cong_name h));
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1295
                  in
80707
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1296
                    (case app_cong () of
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1297
                      NONE => appc ()
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1298
                    | SOME cong =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1299
     (*post processing: some partial applications h t1 ... tj, j <= length ts,
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
  1300
       may be a redex. Example: map (\<lambda>x. x) = (\<lambda>xs. xs) wrt map_cong*)
80707
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1301
                       (let
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1302
                          val thm = congc (prover ctxt) ctxt maxidx cong t0;
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1303
                          val t = the_default t0 (Option.map Thm.rhs_of thm);
80709
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1304
                          val (cl, cr) = Thm.dest_comb t
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1305
                            handle CTERM _ => Thm.dest_comb t0;  (*e.g. congproc has
e6f026505c5b support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
wenzelm
parents: 80707
diff changeset
  1306
                              normalized such that head is removed from t*)
80707
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1307
                          val dVar = Var (("", 0), dummyT);
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1308
                          val skel = list_comb (h, replicate (length ts) dVar);
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1309
                        in
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1310
                          (case botc skel ctxt cl of
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1311
                            NONE => thm
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1312
                          | SOME thm' => transitive3 thm (Thm.combination thm' (Thm.reflexive cr)))
897c993293c5 tuned: anticipate congprocs;
wenzelm
parents: 80706
diff changeset
  1313
                        end handle Pattern.MATCH => appc ()))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1314
                  end)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1315
            | _ => NONE)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1316
        end
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1317
    and impc ct ctxt =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1318
      if mutsimp then mut_impc0 [] ct [] [] ctxt
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1319
      else nonmut_impc ct ctxt
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1320
54984
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1321
    and rules_of_prem prem ctxt =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1322
      if maxidx_of_term (Thm.term_of prem) <> ~1
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1323
      then
55031
wenzelm
parents: 55028
diff changeset
  1324
       (cond_tracing ctxt (fn () =>
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1325
          print_term ctxt "Cannot add premise as rewrite rule because it contains (type) unknowns:"
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1326
            (Thm.term_of prem));
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1327
        (([], NONE), ctxt))
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1328
      else
80713
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
  1329
        let
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
  1330
          val (asm, ctxt') = Thm.assume_hyps prem ctxt;
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
  1331
          val (rews, ctxt'') = extract_safe_rrules asm ctxt';
43e0f32451ee provide Simplifier.set_mksimps_context (roughly following Norbert Schirmer): allow update of context when premises are added to the local simpset;
wenzelm
parents: 80712
diff changeset
  1332
        in ((rews, SOME asm), ctxt'') end
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1333
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1334
    and add_rrules (rrss, asms) ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1335
      (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
  1336
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22902
diff changeset
  1337
    and disch r prem eq =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1338
      let
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1339
        val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq);
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1340
        val eq' =
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1341
          Thm.implies_elim
77879
wenzelm
parents: 77863
diff changeset
  1342
            (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem) (vB, lhs) (vC, rhs))
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74266
diff changeset
  1343
              Drule.imp_cong)
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1344
            (Thm.implies_intr prem eq);
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1345
      in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1346
        if not r then eq'
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1347
        else
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1348
          let
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1349
            val (prem', concl) = Thm.dest_implies lhs;
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1350
            val (prem'', _) = Thm.dest_implies rhs;
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1351
          in
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1352
            Thm.transitive
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1353
              (Thm.transitive
77879
wenzelm
parents: 77863
diff changeset
  1354
                (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem') (vB, prem) (vC, concl))
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74266
diff changeset
  1355
                  Drule.swap_prems_eq)
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1356
                eq')
77879
wenzelm
parents: 77863
diff changeset
  1357
              (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem) (vB, prem'') (vC, concl))
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 74266
diff changeset
  1358
                Drule.swap_prems_eq)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1359
          end
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1360
      end
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1361
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1362
    and rebuild [] _ _ _ _ eq = eq
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1363
      | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ctxt eq =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1364
          let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1365
            val ctxt' = add_rrules (rev rrss, rev asms) ctxt;
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1366
            val concl' =
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1367
              Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq));
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1368
            val dprem = Option.map (disch false prem);
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1369
          in
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1370
            (case rewritec (prover, maxidx) ctxt' concl' of
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1371
              NONE => rebuild prems concl' rrss asms ctxt (dprem eq)
54727
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1372
            | SOME (eq', _) =>
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1373
                transitive2 (fold (disch false) prems (the (transitive3 (dprem eq) eq')))
a806e7251cf0 tuned whitespace;
wenzelm
parents: 54725
diff changeset
  1374
                  (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
  1375
          end
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1376
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1377
    and mut_impc0 prems concl rrss asms ctxt =
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1378
      let
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1379
        val prems' = strip_imp_prems concl;
54984
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1380
        val ((rrss', asms'), ctxt') = fold_map rules_of_prem prems' ctxt |>> split_list;
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1381
      in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1382
        mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss')
54984
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1383
          (asms @ asms') [] [] [] [] ctxt' ~1 ~1
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1384
      end
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1385
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1386
    and mut_impc [] concl [] [] prems' rrss' asms' eqns ctxt changed k =
33245
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33039
diff changeset
  1387
        transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1
65232054ffd0 eliminated some old folds;
wenzelm
parents: 33039
diff changeset
  1388
            (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE)
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1389
          (if changed > 0 then
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1390
             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
  1391
               [] [] [] [] ctxt ~1 changed
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1392
           else rebuild prems' concl rrss' asms' ctxt
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1393
             (botc skel0 (add_rrules (rev rrss', rev asms') ctxt) concl))
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1394
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1395
      | 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
  1396
          prems' rrss' asms' eqns ctxt changed k =
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1397
        (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
  1398
          (rev rrss' @ rrss, rev asms' @ asms) ctxt) prem) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15460
diff changeset
  1399
            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
  1400
              (rrs :: rrss') (asm :: asms') (NONE :: eqns) ctxt changed
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1401
              (if k = 0 then 0 else k - 1)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1402
        | SOME eqn =>
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1403
            let
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 22892
diff changeset
  1404
              val prem' = Thm.rhs_of eqn;
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1405
              val tprems = map Thm.term_of prems;
77863
760515c45864 revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
wenzelm
parents: 77808
diff changeset
  1406
              val i = 1 + fold Integer.max (map (fn p =>
760515c45864 revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
wenzelm
parents: 77808
diff changeset
  1407
                find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1;
54984
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1408
              val ((rrs', asm'), ctxt') = rules_of_prem prem' ctxt;
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1409
            in
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1410
              mut_impc prems concl rrss asms (prem' :: prems')
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1411
                (rrs' :: rrss') (asm' :: asms')
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1412
                (SOME (fold_rev (disch true)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1413
                  (take i prems)
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1414
                  (Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1415
                    (drop i prems, concl))))) :: eqns)
54984
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1416
                ctxt' (length prems') ~1
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1417
            end)
13607
6908230623a3 Completely reimplemented mutual simplification of premises.
berghofe
parents: 13569
diff changeset
  1418
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1419
    (*legacy code -- only for backwards compatibility*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1420
    and nonmut_impc ct ctxt =
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1421
      let
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1422
        val (prem, conc) = Thm.dest_implies ct;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1423
        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
  1424
        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
  1425
        val ctxt1 =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1426
          if not useprem then ctxt
54984
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1427
          else
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1428
            let val ((rrs, asm), ctxt') = rules_of_prem prem1 ctxt
da70ab8531f4 more elementary management of declared hyps, below structure Assumption;
wenzelm
parents: 54982
diff changeset
  1429
            in add_rrules ([rrs], [asm]) ctxt' end;
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1430
      in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1431
        (case botc skel0 ctxt1 conc of
38834
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1432
          NONE =>
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1433
            (case thm1 of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1434
              NONE => NONE
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1435
            | 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
  1436
        | SOME thm2 =>
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1437
            let val thm2' = disch false prem1 thm2 in
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1438
              (case thm1 of
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1439
                NONE => SOME thm2'
658fcba35ed7 more careful treatment of context visibility flag wrt. spurious warnings;
wenzelm
parents: 38715
diff changeset
  1440
              | SOME thm1' =>
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  1441
                 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
  1442
            end)
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1443
      end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1444
54725
fc384e0a7f51 tuned whitespace;
wenzelm
parents: 54724
diff changeset
  1445
  in try_botc end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1446
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1447
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
  1448
(* Meta-rewriting: rewrites t to u and returns the theorem t \<equiv> u *)
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1449
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1450
(*
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1451
  Parameters:
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1452
    mode = (simplify A,
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1453
            use A in simplifying B,
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1454
            use prems of B (if B is again a meta-impl.) to simplify A)
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
  1455
           when simplifying A \<Longrightarrow> B
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1456
    prover: how to solve premises in conditional rewrites and congruences
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1457
*)
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1458
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1459
fun rewrite_cterm mode prover raw_ctxt raw_ct =
17882
b6e44fc46cf0 added set/addloop' for simpset dependent loopers;
wenzelm
parents: 17756
diff changeset
  1460
  let
60324
wenzelm
parents: 60184
diff changeset
  1461
    val ct = raw_ct
78136
wenzelm
parents: 78114
diff changeset
  1462
      |> Thm.transfer_cterm' raw_ctxt
60324
wenzelm
parents: 60184
diff changeset
  1463
      |> Thm.adjust_maxidx_cterm ~1;
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59582
diff changeset
  1464
    val maxidx = Thm.maxidx_of_cterm ct;
52091
9ec2d47de6a7 more rigorous check of simplifier context;
wenzelm
parents: 52082
diff changeset
  1465
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1466
    val ctxt =
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1467
      raw_ctxt
71318
1be996d8bb98 proper proof body context for Simplifier plugins (solvers, loopers, ...) -- avoid crash due to Subgoal.FOCUS (before e58bc223f46c);
wenzelm
parents: 71239
diff changeset
  1468
      |> Variable.set_body true
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1469
      |> Context_Position.set_visible false
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1470
      |> inc_simp_depth
80065
60b6c735b5d5 clarified signature: prefer authentic cterm used in Simplifier, avoid potential re-certification in user-code;
wenzelm
parents: 80064
diff changeset
  1471
      |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, cterm = ct} ctxt);
54729
c5cd7a58cf2d generic trace operations for main steps of Simplifier;
wenzelm
parents: 54728
diff changeset
  1472
55028
00e849f5b397 clarified Simplifier diagnostics -- simplified ML;
wenzelm
parents: 55014
diff changeset
  1473
    val _ =
55031
wenzelm
parents: 55028
diff changeset
  1474
      cond_tracing ctxt (fn () =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1475
        print_term ctxt "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" (Thm.term_of ct));
70472
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
  1476
  in
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
  1477
    ct
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
  1478
    |> bottomc (mode, Option.map (Drule.flexflex_unique (SOME ctxt)) oo prover, maxidx) ctxt
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
  1479
    |> Thm.solve_constraints
cf66d2db97fe more robust and convenient treatment of implicit context;
wenzelm
parents: 69575
diff changeset
  1480
  end;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1481
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1482
val simple_prover =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59058
diff changeset
  1483
  SINGLE o (fn ctxt => ALLGOALS (resolve_tac ctxt (prems_of ctxt)));
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1484
78453
3fdf3c5cfa9d performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents: 78136
diff changeset
  1485
fun rewrite0 ctxt full = rewrite_cterm (full, false, false) simple_prover ctxt;
3fdf3c5cfa9d performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents: 78136
diff changeset
  1486
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1487
fun rewrite _ _ [] = Thm.reflexive
78453
3fdf3c5cfa9d performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents: 78136
diff changeset
  1488
  | rewrite ctxt full thms = rewrite0 (init_simpset thms ctxt) full;
11672
8e75b78f33f3 full_rewrite_cterm_aux (see also tactic.ML);
wenzelm
parents: 11505
diff changeset
  1489
78453
3fdf3c5cfa9d performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
wenzelm
parents: 78136
diff changeset
  1490
fun rewrite0_rule ctxt = Conv.fconv_rule (rewrite0 ctxt true);
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1491
fun rewrite_rule ctxt = Conv.fconv_rule o rewrite ctxt true;
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1492
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1493
(*simple term rewriting -- no proof*)
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16380
diff changeset
  1494
fun rewrite_term thy rules procs =
17203
29b2563f5c11 refer to theory instead of low-level tsig;
wenzelm
parents: 17046
diff changeset
  1495
  Pattern.rewrite_term thy (map decomp_simp' rules) procs;
15023
0e4689f411d5 major cleanup; got rid of obsolete meta_simpset;
wenzelm
parents: 15011
diff changeset
  1496
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1497
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
  1498
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
  1499
(*Rewrite the subgoals of a proof state (represented by a theorem)*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1500
fun rewrite_goals_rule ctxt thms th =
23584
9b5ba76de1c2 tuned goal conversion interfaces;
wenzelm
parents: 23536
diff changeset
  1501
  Conv.fconv_rule (Conv.prems_conv ~1 (rewrite_cterm (true, true, true) simple_prover
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1502
    (init_simpset thms ctxt))) th;
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1503
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1504
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1505
(** meta-rewriting tactics **)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1506
28839
32d498cf7595 eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
wenzelm
parents: 28620
diff changeset
  1507
(*Rewrite all subgoals*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1508
fun rewrite_goals_tac ctxt defs = PRIMITIVE (rewrite_goals_rule ctxt defs);
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1509
28839
32d498cf7595 eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
wenzelm
parents: 28620
diff changeset
  1510
(*Rewrite one subgoal*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1511
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
  1512
  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
  1513
    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
  1514
  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
  1515
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1516
fun rewrite_goal_tac ctxt thms =
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1517
  generic_rewrite_goal_tac (true, false, false) (K no_tac) (init_simpset thms ctxt);
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
  1518
46707
1427dcc7c9a6 eliminated odd comment from distant past;
wenzelm
parents: 46465
diff changeset
  1519
(*Prunes all redundant parameters from the proof state by rewriting.*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1520
fun prune_params_tac ctxt = rewrite_goals_tac ctxt [Drule.triv_forall_equality];
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1521
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1522
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1523
(* for folding definitions, handling critical pairs *)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1524
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1525
(*The depth of nesting in a term*)
32797
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
  1526
fun term_depth (Abs (_, _, t)) = 1 + term_depth t
2e8fae2d998c eliminated dead code;
wenzelm
parents: 32738
diff changeset
  1527
  | 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
  1528
  | term_depth _ = 0;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1529
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
  1530
val lhs_of_thm = #1 o Logic.dest_equals o Thm.prop_of;
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1531
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
  1532
(*folding should handle critical pairs!  E.g. K \<equiv> Inl 0,  S \<equiv> Inr (Inl 0)
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1533
  Returns longest lhs first to avoid folding its subexpressions.*)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1534
fun sort_lhs_depths defs =
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1535
  let val keylist = AList.make (term_depth o lhs_of_thm) defs
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1536
      val keys = sort_distinct (rev_order o int_ord) (map #2 keylist)
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1537
  in map (AList.find (op =) keylist) keys end;
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1538
36944
dbf831a50e4a less pervasive names from structure Thm;
wenzelm
parents: 36610
diff changeset
  1539
val rev_defs = sort_lhs_depths o map Thm.symmetric;
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1540
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1541
fun fold_rule ctxt defs = fold (rewrite_rule ctxt) (rev_defs defs);
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54731
diff changeset
  1542
fun fold_goals_tac ctxt defs = EVERY (map (rewrite_goals_tac ctxt) (rev_defs defs));
21708
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1543
45e7491bea47 reorganized structure Tactic vs. MetaSimplifier;
wenzelm
parents: 21646
diff changeset
  1544
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67649
diff changeset
  1545
(* HHF normal form: \<And> before \<Longrightarrow>, outermost \<And> generalized *)
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1546
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1547
local
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1548
78046
78deba4fdf27 tuned: more accurate check (is_norm_hhf protect);
wenzelm
parents: 77879
diff changeset
  1549
fun gen_norm_hhf protect ss ctxt0 th0 =
71177
71467e35fc3c more structural integrity;
wenzelm
parents: 70586
diff changeset
  1550
  let
71467e35fc3c more structural integrity;
wenzelm
parents: 70586
diff changeset
  1551
    val (ctxt, th) = Thm.join_transfer_context (ctxt0, th0);
71467e35fc3c more structural integrity;
wenzelm
parents: 70586
diff changeset
  1552
    val th' =
78046
78deba4fdf27 tuned: more accurate check (is_norm_hhf protect);
wenzelm
parents: 77879
diff changeset
  1553
      if Drule.is_norm_hhf protect (Thm.prop_of th) then th
71177
71467e35fc3c more structural integrity;
wenzelm
parents: 70586
diff changeset
  1554
      else
71467e35fc3c more structural integrity;
wenzelm
parents: 70586
diff changeset
  1555
        Conv.fconv_rule (rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset ss ctxt)) th;
71467e35fc3c more structural integrity;
wenzelm
parents: 70586
diff changeset
  1556
  in th' |> Thm.adjust_maxidx_thm ~1 |> Variable.gen_all ctxt end;
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1557
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1558
val hhf_ss =
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1559
  Context.the_local_context ()
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1560
  |> init_simpset Drule.norm_hhf_eqs
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1561
  |> simpset_of;
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1562
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51611
diff changeset
  1563
val hhf_protect_ss =
63221
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1564
  Context.the_local_context ()
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1565
  |> init_simpset Drule.norm_hhf_eqs
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1566
  |> add_eqcong Drule.protect_cong
7d43fbbaba28 avoid warnings on duplicate rules in the given list;
wenzelm
parents: 62913
diff changeset
  1567
  |> simpset_of;
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1568
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1569
in
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1570
78046
78deba4fdf27 tuned: more accurate check (is_norm_hhf protect);
wenzelm
parents: 77879
diff changeset
  1571
val norm_hhf = gen_norm_hhf {protect = false} hhf_ss;
78deba4fdf27 tuned: more accurate check (is_norm_hhf protect);
wenzelm
parents: 77879
diff changeset
  1572
val norm_hhf_protect = gen_norm_hhf {protect = true} hhf_protect_ss;
20228
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1573
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1574
end;
e0f9e8a6556b moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
wenzelm
parents: 20197
diff changeset
  1575
10413
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1576
end;
0e015d9bea4e Added new file meta_simplifier.ML
berghofe
parents:
diff changeset
  1577
41228
e1fce873b814 renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents: 41227
diff changeset
  1578
structure Basic_Meta_Simplifier: BASIC_RAW_SIMPLIFIER = Raw_Simplifier;
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 31298
diff changeset
  1579
open Basic_Meta_Simplifier;