src/Provers/classical.ML
author immler
Sun, 27 Oct 2019 21:51:14 -0400
changeset 71035 6fe5a0e1fa8e
parent 69593 3dda49e08b9d
child 74561 8e6c973003c8
permissions -rw-r--r--
moved theory Interval from the AFP
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9938
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
     1
(*  Title:      Provers/classical.ML
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
Theorem prover for classical reasoning, including predicate calculus, set
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
theory, etc.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
     7
Rules must be classified as intro, elim, safe, unsafe.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
A rule is unsafe unless it can be applied blindly without harmful results.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
For a rule to be safe, its premises and conclusion should be logically
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
equivalent.  There should be no variables in the premises that are not in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
the conclusion.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    15
(*higher precedence than := facilitates use of references*)
12376
480303e3fa08 simplified (and clarified) integration with Pure/ContextRules;
wenzelm
parents: 12362
diff changeset
    16
infix 4 addSIs addSEs addSDs addIs addEs addDs delrules
4651
70dd492a1698 changed wrapper mechanism of classical reasoner
oheimb
parents: 4646
diff changeset
    17
  addSWrapper delSWrapper addWrapper delWrapper
11181
d04f57b91166 renamed addaltern to addafter, addSaltern to addSafter
oheimb
parents: 11168
diff changeset
    18
  addSbefore addSafter addbefore addafter
5523
dc8cdc192cd0 added addD2, addE2, addSD2, and addSE2
oheimb
parents: 5028
diff changeset
    19
  addD2 addE2 addSD2 addSE2;
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    20
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
signature CLASSICAL_DATA =
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    22
sig
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    23
  val imp_elim: thm  (* P --> Q ==> (~ R ==> P) ==> (Q ==> R) ==> R *)
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    24
  val not_elim: thm  (* ~P ==> P ==> R *)
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    25
  val swap: thm  (* ~ P ==> (~ R ==> P) ==> R *)
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    26
  val classical: thm  (* (~ P ==> P) ==> P *)
50062
e038198f7d08 more concise/precise documentation;
wenzelm
parents: 48126
diff changeset
    27
  val sizef: thm -> int  (* size function for BEST_FIRST, typically size_of_thm *)
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51718
diff changeset
    28
  val hyp_subst_tacs: (Proof.context -> int -> tactic) list (* optional tactics for
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51718
diff changeset
    29
    substitution in the hypotheses; assumed to be safe! *)
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    30
end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
    32
signature BASIC_CLASSICAL =
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    33
sig
42812
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
    34
  type wrapper = (int -> tactic) -> int -> tactic
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
  type claset
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    36
  val print_claset: Proof.context -> unit
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    37
  val addDs: Proof.context * thm list -> Proof.context
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    38
  val addEs: Proof.context * thm list -> Proof.context
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    39
  val addIs: Proof.context * thm list -> Proof.context
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    40
  val addSDs: Proof.context * thm list -> Proof.context
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    41
  val addSEs: Proof.context * thm list -> Proof.context
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    42
  val addSIs: Proof.context * thm list -> Proof.context
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    43
  val delrules: Proof.context * thm list -> Proof.context
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    44
  val addSWrapper: Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    45
  val delSWrapper: Proof.context * string -> Proof.context
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    46
  val addWrapper: Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    47
  val delWrapper: Proof.context * string -> Proof.context
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
    48
  val addSbefore: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
    49
  val addSafter: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
    50
  val addbefore: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
    51
  val addafter: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    52
  val addD2: Proof.context * (string * thm) -> Proof.context
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    53
  val addE2: Proof.context * (string * thm) -> Proof.context
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    54
  val addSD2: Proof.context * (string * thm) -> Proof.context
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    55
  val addSE2: Proof.context * (string * thm) -> Proof.context
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    56
  val appSWrappers: Proof.context -> wrapper
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    57
  val appWrappers: Proof.context -> wrapper
982
4fe0b642b7d5 Addition of wrappers for integration with the simplifier.
lcp
parents: 747
diff changeset
    58
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    59
  val claset_of: Proof.context -> claset
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    60
  val put_claset: claset -> Proof.context -> Proof.context
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    61
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    62
  val map_theory_claset: (Proof.context -> Proof.context) -> theory -> theory
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
    63
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    64
  val fast_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    65
  val slow_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    66
  val astar_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    67
  val slow_astar_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    68
  val best_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    69
  val first_best_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    70
  val slow_best_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    71
  val depth_tac: Proof.context -> int -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    72
  val deepen_tac: Proof.context -> int -> int -> tactic
1587
e7d8a4957bac Now provides astar versions (thanks to Norbert Voelker)
paulson
parents: 1524
diff changeset
    73
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
    74
  val contr_tac: Proof.context -> int -> tactic
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
    75
  val dup_elim: Proof.context -> thm -> thm
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    76
  val dup_intr: thm -> thm
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    77
  val dup_step_tac: Proof.context -> int -> tactic
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
    78
  val eq_mp_tac: Proof.context -> int -> tactic
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
    79
  val unsafe_step_tac: Proof.context -> int -> tactic
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
    80
  val mp_tac: Proof.context -> int -> tactic
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    81
  val safe_tac: Proof.context -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    82
  val safe_steps_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    83
  val safe_step_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    84
  val clarify_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    85
  val clarify_step_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    86
  val step_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    87
  val slow_step_tac: Proof.context -> int -> tactic
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
    88
  val swapify: thm list -> thm list
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
    89
  val swap_res_tac: Proof.context -> thm list -> int -> tactic
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    90
  val inst_step_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    91
  val inst0_step_tac: Proof.context -> int -> tactic
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
    92
  val instp_step_tac: Proof.context -> int -> tactic
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
    93
end;
1724
bb02e6976258 Added functions for default claset.
berghofe
parents: 1711
diff changeset
    94
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
    95
signature CLASSICAL =
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
    96
sig
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
    97
  include BASIC_CLASSICAL
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
    98
  val classical_rule: Proof.context -> thm -> thm
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
    99
  type rule = thm * (thm * thm list) * (thm * thm list)
42812
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   100
  type netpair = (int * (bool * thm)) Net.net * (int * (bool * thm)) Net.net
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   101
  val rep_cs: claset ->
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   102
   {safeIs: rule Item_Net.T,
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   103
    safeEs: rule Item_Net.T,
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   104
    unsafeIs: rule Item_Net.T,
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   105
    unsafeEs: rule Item_Net.T,
42812
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   106
    swrappers: (string * (Proof.context -> wrapper)) list,
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   107
    uwrappers: (string * (Proof.context -> wrapper)) list,
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   108
    safe0_netpair: netpair,
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   109
    safep_netpair: netpair,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   110
    unsafe_netpair: netpair,
42812
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   111
    dup_netpair: netpair,
60942
wenzelm
parents: 60817
diff changeset
   112
    extra_netpair: Context_Rules.netpair}
24021
491c68f40bc4 added get_cs/map_cs;
wenzelm
parents: 23594
diff changeset
   113
  val get_cs: Context.generic -> claset
491c68f40bc4 added get_cs/map_cs;
wenzelm
parents: 23594
diff changeset
   114
  val map_cs: (claset -> claset) -> Context.generic -> Context.generic
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   115
  val safe_dest: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   116
  val safe_elim: int option -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   117
  val safe_intro: int option -> attribute
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   118
  val unsafe_dest: int option -> attribute
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   119
  val unsafe_elim: int option -> attribute
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   120
  val unsafe_intro: int option -> attribute
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   121
  val rule_del: attribute
61327
0a4c364df431 tuned signature;
wenzelm
parents: 61268
diff changeset
   122
  val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic
0a4c364df431 tuned signature;
wenzelm
parents: 61268
diff changeset
   123
  val standard_tac: Proof.context -> thm list -> tactic
30513
1796b8ea88aa eliminated type Args.T;
wenzelm
parents: 30510
diff changeset
   124
  val cla_modifiers: Method.modifier parser list
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   125
  val cla_method:
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   126
    (Proof.context -> tactic) -> (Proof.context -> Proof.method) context_parser
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   127
  val cla_method':
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   128
    (Proof.context -> int -> tactic) -> (Proof.context -> Proof.method) context_parser
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   129
end;
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   130
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
42799
4e33894aec6d modernized functor names;
wenzelm
parents: 42798
diff changeset
   132
functor Classical(Data: CLASSICAL_DATA): CLASSICAL =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   135
(** classical elimination rules **)
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   136
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   137
(*
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   138
Classical reasoning requires stronger elimination rules.  For
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   139
instance, make_elim of Pure transforms the HOL rule injD into
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   140
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   141
    [| inj f; f x = f y; x = y ==> PROP W |] ==> PROP W
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   142
26938
64e850c3da9e tuned comments;
wenzelm
parents: 26928
diff changeset
   143
Such rules can cause fast_tac to fail and blast_tac to report "PROOF
59119
c90c02940964 tuned spelling;
wenzelm
parents: 58963
diff changeset
   144
FAILED"; classical_rule will strengthen this to
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   145
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   146
    [| inj f; ~ W ==> f x = f y; x = y ==> W |] ==> W
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   147
*)
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   148
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   149
fun classical_rule ctxt rule =
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   150
  if is_some (Object_Logic.elim_concl ctxt rule) then
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   151
    let
60817
3f38ed5a02c1 clarified context;
wenzelm
parents: 60757
diff changeset
   152
      val thy = Proof_Context.theory_of ctxt;
42792
85fb70b0c5e8 do not open ML structures;
wenzelm
parents: 42791
diff changeset
   153
      val rule' = rule RS Data.classical;
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   154
      val concl' = Thm.concl_of rule';
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   155
      fun redundant_hyp goal =
19257
4463aee061bc ObjectLogic.is_elim;
wenzelm
parents: 19110
diff changeset
   156
        concl' aconv Logic.strip_assums_concl goal orelse
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   157
          (case Logic.strip_assums_hyp goal of
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   158
            hyp :: hyps => exists (fn t => t aconv hyp) hyps
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   159
          | _ => false);
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   160
      val rule'' =
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   161
        rule' |> ALLGOALS (SUBGOAL (fn (goal, i) =>
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   162
          if i = 1 orelse redundant_hyp goal
60757
c09598a97436 prefer tactics with explicit context;
wenzelm
parents: 60650
diff changeset
   163
          then eresolve_tac ctxt [thin_rl] i
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   164
          else all_tac))
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   165
        |> Seq.hd
21963
416a5338d2bb removed obsolete name_hint handling;
wenzelm
parents: 21689
diff changeset
   166
        |> Drule.zero_var_indexes;
60817
3f38ed5a02c1 clarified context;
wenzelm
parents: 60757
diff changeset
   167
    in if Thm.equiv_thm thy (rule, rule'') then rule else rule'' end
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   168
  else rule;
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   169
23594
e65e466dda01 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23178
diff changeset
   170
(*flatten nested meta connectives in prems*)
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   171
fun flat_rule ctxt =
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   172
  Conv.fconv_rule (Conv.prems_conv ~1 (Object_Logic.atomize_prems ctxt));
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   173
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   174
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   175
(*** Useful tactics for classical reasoning ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   177
(*Prove goal that assumes both P and ~P.
4392
ea41d9c1b0ef More deterministic (?) contr_tac
paulson
parents: 4380
diff changeset
   178
  No backtracking if it finds an equal assumption.  Perhaps should call
ea41d9c1b0ef More deterministic (?) contr_tac
paulson
parents: 4380
diff changeset
   179
  ematch_tac instead of eresolve_tac, but then cannot prove ZF/cantor.*)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
   180
fun contr_tac ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   181
  eresolve_tac ctxt [Data.not_elim] THEN' (eq_assume_tac ORELSE' assume_tac ctxt);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   182
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   183
(*Finds P-->Q and P in the assumptions, replaces implication by Q.
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   184
  Could do the same thing for P<->Q and P... *)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   185
fun mp_tac ctxt i =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   186
  eresolve_tac ctxt [Data.not_elim, Data.imp_elim] i THEN assume_tac ctxt i;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
(*Like mp_tac but instantiates no variables*)
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   189
fun eq_mp_tac ctxt i = ematch_tac ctxt [Data.not_elim, Data.imp_elim] i THEN eq_assume_tac i;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   191
(*Creates rules to eliminate ~A, from rules to introduce A*)
26412
0918f5c0bbca pass imp_elim (instead of mp) and swap explicitly -- avoids store_thm;
wenzelm
parents: 24867
diff changeset
   192
fun swapify intrs = intrs RLN (2, [Data.swap]);
61853
fb7756087101 rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
wenzelm
parents: 61327
diff changeset
   193
val swapped = Thm.rule_attribute [] (fn _ => fn th => th RSN (2, Data.swap));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   195
(*Uses introduction rules in the normal way, or on negated assumptions,
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   196
  trying rules in order. *)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
   197
fun swap_res_tac ctxt rls =
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   198
  let
67649
1e1782c1aedf tuned signature;
wenzelm
parents: 64556
diff changeset
   199
    val transfer = Thm.transfer' ctxt;
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   200
    fun addrl rl brls = (false, transfer rl) :: (true, transfer rl RSN (2, Data.swap)) :: brls;
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   201
  in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
   202
    assume_tac ctxt ORELSE'
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
   203
    contr_tac ctxt ORELSE'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   204
    biresolve_tac ctxt (fold_rev addrl rls [])
42792
85fb70b0c5e8 do not open ML structures;
wenzelm
parents: 42791
diff changeset
   205
  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   206
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   207
(*Duplication of unsafe rules, for complete provers*)
42792
85fb70b0c5e8 do not open ML structures;
wenzelm
parents: 42791
diff changeset
   208
fun dup_intr th = zero_var_indexes (th RS Data.classical);
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   209
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   210
fun dup_elim ctxt th =
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   211
  let val rl = (th RSN (2, revcut_rl)) |> Thm.assumption (SOME ctxt) 2 |> Seq.hd;
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   212
  in rule_by_tactic ctxt (TRYALL (eresolve_tac ctxt [revcut_rl])) rl end;
36546
a9873318fe30 proper context for rule_by_tactic;
wenzelm
parents: 35625
diff changeset
   213
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   214
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   215
(**** Classical rule sets ****)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   217
type rule = thm * (thm * thm list) * (thm * thm list);
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   218
  (*external form, internal form (possibly swapped), dup form (possibly swapped)*)
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   219
42812
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   220
type netpair = (int * (bool * thm)) Net.net * (int * (bool * thm)) Net.net;
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   221
type wrapper = (int -> tactic) -> int -> tactic;
dda4aef7cba4 tuned signature;
wenzelm
parents: 42810
diff changeset
   222
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   223
datatype claset =
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   224
  CS of
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   225
   {safeIs: rule Item_Net.T,  (*safe introduction rules*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   226
    safeEs: rule Item_Net.T,  (*safe elimination rules*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   227
    unsafeIs: rule Item_Net.T,  (*unsafe introduction rules*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   228
    unsafeEs: rule Item_Net.T,  (*unsafe elimination rules*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   229
    swrappers: (string * (Proof.context -> wrapper)) list,  (*for transforming safe_step_tac*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   230
    uwrappers: (string * (Proof.context -> wrapper)) list,  (*for transforming step_tac*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   231
    safe0_netpair: netpair,  (*nets for trivial cases*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   232
    safep_netpair: netpair,  (*nets for >0 subgoals*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   233
    unsafe_netpair: netpair,  (*nets for unsafe rules*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   234
    dup_netpair: netpair,  (*nets for duplication*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   235
    extra_netpair: Context_Rules.netpair};  (*nets for extra rules*)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   236
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   237
val empty_rules: rule Item_Net.T =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   238
  Item_Net.init (Thm.eq_thm_prop o apply2 #1) (single o Thm.full_prop_of o #1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   239
6502
bc30e13b36a8 improved 'single' method;
wenzelm
parents: 6492
diff changeset
   240
val empty_netpair = (Net.empty, Net.empty);
bc30e13b36a8 improved 'single' method;
wenzelm
parents: 6492
diff changeset
   241
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   242
val empty_cs =
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   243
  CS
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   244
   {safeIs = empty_rules,
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   245
    safeEs = empty_rules,
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   246
    unsafeIs = empty_rules,
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   247
    unsafeEs = empty_rules,
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   248
    swrappers = [],
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   249
    uwrappers = [],
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   250
    safe0_netpair = empty_netpair,
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   251
    safep_netpair = empty_netpair,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   252
    unsafe_netpair = empty_netpair,
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   253
    dup_netpair = empty_netpair,
60942
wenzelm
parents: 60817
diff changeset
   254
    extra_netpair = empty_netpair};
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   255
4653
d60f76680bf4 renamed rep_claset to rep_cs
oheimb
parents: 4651
diff changeset
   256
fun rep_cs (CS args) = args;
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   257
4079
9df5e4f22d96 new implicit claset mechanism based on Sign.sg anytype data;
wenzelm
parents: 4066
diff changeset
   258
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   259
(*** Adding (un)safe introduction or elimination rules.
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   260
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   261
    In case of overlap, new rules are tried BEFORE old ones!!
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   262
***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   263
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   264
fun joinrules (intrs, elims) = map (pair true) elims @ map (pair false) intrs;
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   265
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   266
(*Priority: prefer rules with fewest subgoals,
1231
91d2c1bb5803 clarified comment
paulson
parents: 1073
diff changeset
   267
  then rules added most recently (preferring the head of the list).*)
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   268
fun tag_brls k [] = []
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   269
  | tag_brls k (brl::brls) =
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   270
      (1000000*subgoals_of_brl brl + k, brl) ::
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   271
      tag_brls (k+1) brls;
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   272
12401
4363432ef0cd added 'swapped' attribute;
wenzelm
parents: 12376
diff changeset
   273
fun tag_brls' _ _ [] = []
4363432ef0cd added 'swapped' attribute;
wenzelm
parents: 12376
diff changeset
   274
  | tag_brls' w k (brl::brls) = ((w, k), brl) :: tag_brls' w (k + 1) brls;
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   275
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22846
diff changeset
   276
fun insert_tagged_list rls = fold_rev Tactic.insert_tagged_brl rls;
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   277
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   278
(*Insert into netpair that already has nI intr rules and nE elim rules.
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   279
  Count the intr rules double (to account for swapify).  Negate to give the
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   280
  new insertions the lowest priority.*)
12376
480303e3fa08 simplified (and clarified) integration with Pure/ContextRules;
wenzelm
parents: 12362
diff changeset
   281
fun insert (nI, nE) = insert_tagged_list o (tag_brls (~(2*nI+nE))) o joinrules;
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   282
fun insert' w (nI, nE) = insert_tagged_list o tag_brls' w (~(nI + nE)) o joinrules;
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   283
23178
07ba6b58b3d2 simplified/unified list fold;
wenzelm
parents: 22846
diff changeset
   284
fun delete_tagged_list rls = fold_rev Tactic.delete_tagged_brl rls;
12362
57cd572103c4 made SML/NJ happy;
wenzelm
parents: 12311
diff changeset
   285
fun delete x = delete_tagged_list (joinrules x);
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   286
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61056
diff changeset
   287
fun bad_thm ctxt msg th = error (msg ^ "\n" ^ Thm.string_of_thm ctxt th);
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   288
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   289
fun make_elim ctxt th =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   290
  if has_fewer_prems 1 th then bad_thm ctxt "Ill-formed destruction rule" th
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   291
  else Tactic.make_elim th;
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   292
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   293
fun warn_thm ctxt msg th =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   294
  if Context_Position.is_really_visible ctxt
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61056
diff changeset
   295
  then warning (msg ^ Thm.string_of_thm ctxt th) else ();
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   296
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   297
fun warn_rules ctxt msg rules (r: rule) =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   298
  Item_Net.member rules r andalso (warn_thm ctxt msg (#1 r); true);
42807
e639d91d9073 more precise warnings: observe context visibility;
wenzelm
parents: 42799
diff changeset
   299
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   300
fun warn_claset ctxt r (CS {safeIs, safeEs, unsafeIs, unsafeEs, ...}) =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   301
  warn_rules ctxt "Rule already declared as safe introduction (intro!)\n" safeIs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   302
  warn_rules ctxt "Rule already declared as safe elimination (elim!)\n" safeEs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   303
  warn_rules ctxt "Rule already declared as introduction (intro)\n" unsafeIs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   304
  warn_rules ctxt "Rule already declared as elimination (elim)\n" unsafeEs r;
1927
6f97cb16e453 New classical reasoner: warns of, and ignores, redundant rules.
paulson
parents: 1814
diff changeset
   305
12376
480303e3fa08 simplified (and clarified) integration with Pure/ContextRules;
wenzelm
parents: 12362
diff changeset
   306
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   307
(*** add rules ***)
982
4fe0b642b7d5 Addition of wrappers for integration with the simplifier.
lcp
parents: 747
diff changeset
   308
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   309
fun add_safe_intro w r
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   310
    (cs as CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   311
      safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   312
  if Item_Net.member safeIs r then cs
1927
6f97cb16e453 New classical reasoner: warns of, and ignores, redundant rules.
paulson
parents: 1814
diff changeset
   313
  else
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   314
    let
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   315
      val (th, rl, _) = r;
23594
e65e466dda01 renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
wenzelm
parents: 23178
diff changeset
   316
      val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   317
        List.partition (Thm.no_prems o fst) [rl];
42810
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   318
      val nI = Item_Net.length safeIs + 1;
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   319
      val nE = Item_Net.length safeEs;
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   320
    in
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   321
      CS
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   322
       {safeIs = Item_Net.update r safeIs,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   323
        safe0_netpair = insert (nI, nE) (map fst safe0_rls, maps snd safe0_rls) safe0_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   324
        safep_netpair = insert (nI, nE) (map fst safep_rls, maps snd safep_rls) safep_netpair,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   325
        safeEs = safeEs,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   326
        unsafeIs = unsafeIs,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   327
        unsafeEs = unsafeEs,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   328
        swrappers = swrappers,
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   329
        uwrappers = uwrappers,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   330
        unsafe_netpair = unsafe_netpair,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   331
        dup_netpair = dup_netpair,
60944
bb75b61dba5d tuned whitespace;
wenzelm
parents: 60943
diff changeset
   332
        extra_netpair = insert' (the_default 0 w) (nI, nE) ([th], []) extra_netpair}
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   333
    end;
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   334
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   335
fun add_safe_elim w r
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   336
    (cs as CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   337
      safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   338
  if Item_Net.member safeEs r then cs
1927
6f97cb16e453 New classical reasoner: warns of, and ignores, redundant rules.
paulson
parents: 1814
diff changeset
   339
  else
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   340
    let
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   341
      val (th, rl, _) = r;
18534
6799b38ed872 added classical_rule, which replaces Data.make_elim;
wenzelm
parents: 18374
diff changeset
   342
      val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*)
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   343
        List.partition (fn (rl, _) => Thm.nprems_of rl = 1) [rl];
42810
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   344
      val nI = Item_Net.length safeIs;
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   345
      val nE = Item_Net.length safeEs + 1;
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   346
    in
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   347
      CS
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   348
       {safeEs = Item_Net.update r safeEs,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   349
        safe0_netpair = insert (nI, nE) ([], map fst safe0_rls) safe0_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   350
        safep_netpair = insert (nI, nE) ([], map fst safep_rls) safep_netpair,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   351
        safeIs = safeIs,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   352
        unsafeIs = unsafeIs,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   353
        unsafeEs = unsafeEs,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   354
        swrappers = swrappers,
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   355
        uwrappers = uwrappers,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   356
        unsafe_netpair = unsafe_netpair,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   357
        dup_netpair = dup_netpair,
60944
bb75b61dba5d tuned whitespace;
wenzelm
parents: 60943
diff changeset
   358
        extra_netpair = insert' (the_default 0 w) (nI, nE) ([], [th]) extra_netpair}
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   359
    end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   360
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   361
fun add_unsafe_intro w r
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   362
    (cs as CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   363
      safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   364
  if Item_Net.member unsafeIs r then cs
1927
6f97cb16e453 New classical reasoner: warns of, and ignores, redundant rules.
paulson
parents: 1814
diff changeset
   365
  else
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   366
    let
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   367
      val (th, rl, dup_rl) = r;
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   368
      val nI = Item_Net.length unsafeIs + 1;
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   369
      val nE = Item_Net.length unsafeEs;
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   370
    in
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   371
      CS
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   372
       {unsafeIs = Item_Net.update r unsafeIs,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   373
        unsafe_netpair = insert (nI, nE) ([fst rl], snd rl) unsafe_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   374
        dup_netpair = insert (nI, nE) ([fst dup_rl], snd dup_rl) dup_netpair,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   375
        safeIs = safeIs,
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   376
        safeEs = safeEs,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   377
        unsafeEs = unsafeEs,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   378
        swrappers = swrappers,
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   379
        uwrappers = uwrappers,
9938
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
   380
        safe0_netpair = safe0_netpair,
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
   381
        safep_netpair = safep_netpair,
60942
wenzelm
parents: 60817
diff changeset
   382
        extra_netpair = insert' (the_default 1 w) (nI, nE) ([th], []) extra_netpair}
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   383
    end;
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   384
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   385
fun add_unsafe_elim w r
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   386
    (cs as CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   387
      safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   388
  if Item_Net.member unsafeEs r then cs
1927
6f97cb16e453 New classical reasoner: warns of, and ignores, redundant rules.
paulson
parents: 1814
diff changeset
   389
  else
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   390
    let
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   391
      val (th, rl, dup_rl) = r;
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   392
      val nI = Item_Net.length unsafeIs;
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   393
      val nE = Item_Net.length unsafeEs + 1;
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   394
    in
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   395
      CS
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   396
       {unsafeEs = Item_Net.update r unsafeEs,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   397
        unsafe_netpair = insert (nI, nE) ([], [fst rl]) unsafe_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   398
        dup_netpair = insert (nI, nE) ([], [fst dup_rl]) dup_netpair,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   399
        safeIs = safeIs,
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   400
        safeEs = safeEs,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   401
        unsafeIs = unsafeIs,
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   402
        swrappers = swrappers,
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   403
        uwrappers = uwrappers,
9938
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
   404
        safe0_netpair = safe0_netpair,
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
   405
        safep_netpair = safep_netpair,
60942
wenzelm
parents: 60817
diff changeset
   406
        extra_netpair = insert' (the_default 1 w) (nI, nE) ([], [th]) extra_netpair}
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   407
    end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   408
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   409
fun trim_context (th, (th1, ths1), (th2, ths2)) =
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   410
  (Thm.trim_context th,
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   411
    (Thm.trim_context th1, map Thm.trim_context ths1),
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   412
    (Thm.trim_context th2, map Thm.trim_context ths2));
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   413
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   414
fun addSI w ctxt th (cs as CS {safeIs, ...}) =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   415
  let
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   416
    val th' = flat_rule ctxt th;
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   417
    val rl = (th', swapify [th']);
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   418
    val r = trim_context (th, rl, rl);
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   419
    val _ =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   420
      warn_rules ctxt "Ignoring duplicate safe introduction (intro!)\n" safeIs r orelse
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   421
      warn_claset ctxt r cs;
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   422
  in add_safe_intro w r cs end;
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   423
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   424
fun addSE w ctxt th (cs as CS {safeEs, ...}) =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   425
  let
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   426
    val _ = has_fewer_prems 1 th andalso bad_thm ctxt "Ill-formed elimination rule" th;
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   427
    val th' = classical_rule ctxt (flat_rule ctxt th);
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   428
    val rl = (th', []);
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   429
    val r = trim_context (th, rl, rl);
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   430
    val _ =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   431
      warn_rules ctxt "Ignoring duplicate safe elimination (elim!)\n" safeEs r orelse
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   432
      warn_claset ctxt r cs;
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   433
  in add_safe_elim w r cs end;
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   434
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   435
fun addSD w ctxt th = addSE w ctxt (make_elim ctxt th);
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   436
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   437
fun addI w ctxt th (cs as CS {unsafeIs, ...}) =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   438
  let
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   439
    val th' = flat_rule ctxt th;
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   440
    val dup_th' = dup_intr th' handle THM _ => bad_thm ctxt "Ill-formed introduction rule" th;
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   441
    val r = trim_context (th, (th', swapify [th']), (dup_th', swapify [dup_th']));
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   442
    val _ =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   443
      warn_rules ctxt "Ignoring duplicate introduction (intro)\n" unsafeIs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   444
      warn_claset ctxt r cs;
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   445
  in add_unsafe_intro w r cs end;
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   446
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   447
fun addE w ctxt th (cs as CS {unsafeEs, ...}) =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   448
  let
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   449
    val _ = has_fewer_prems 1 th andalso bad_thm ctxt "Ill-formed elimination rule" th;
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   450
    val th' = classical_rule ctxt (flat_rule ctxt th);
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   451
    val dup_th' = dup_elim ctxt th' handle THM _ => bad_thm ctxt "Ill-formed elimination rule" th;
61056
e9d08b88d2cc trim context for persistent storage;
wenzelm
parents: 61055
diff changeset
   452
    val r = trim_context (th, (th', []), (dup_th', []));
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   453
    val _ =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   454
      warn_rules ctxt "Ignoring duplicate elimination (elim)\n" unsafeEs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   455
      warn_claset ctxt r cs;
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   456
  in add_unsafe_elim w r cs end;
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   457
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   458
fun addD w ctxt th = addE w ctxt (make_elim ctxt th);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   459
1073
b3f190995bc9 Recoded addSIs, etc., so that nets are built incrementally
lcp
parents: 1010
diff changeset
   460
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   461
(*** delete rules ***)
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   462
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   463
local
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   464
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   465
fun del_safe_intro (r: rule)
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   466
  (CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   467
    safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   468
  let
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   469
    val (th, rl, _) = r;
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   470
    val (safe0_rls, safep_rls) = List.partition (Thm.no_prems o fst) [rl];
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   471
  in
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   472
    CS
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   473
     {safe0_netpair = delete (map fst safe0_rls, maps snd safe0_rls) safe0_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   474
      safep_netpair = delete (map fst safep_rls, maps snd safep_rls) safep_netpair,
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   475
      safeIs = Item_Net.remove r safeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   476
      safeEs = safeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   477
      unsafeIs = unsafeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   478
      unsafeEs = unsafeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   479
      swrappers = swrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   480
      uwrappers = uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   481
      unsafe_netpair = unsafe_netpair,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   482
      dup_netpair = dup_netpair,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   483
      extra_netpair = delete ([th], []) extra_netpair}
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   484
  end;
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   485
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   486
fun del_safe_elim (r: rule)
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   487
  (CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   488
    safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   489
  let
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   490
    val (th, rl, _) = r;
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   491
    val (safe0_rls, safep_rls) = List.partition (fn (rl, _) => Thm.nprems_of rl = 1) [rl];
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   492
  in
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   493
    CS
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   494
     {safe0_netpair = delete ([], map fst safe0_rls) safe0_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   495
      safep_netpair = delete ([], map fst safep_rls) safep_netpair,
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   496
      safeIs = safeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   497
      safeEs = Item_Net.remove r safeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   498
      unsafeIs = unsafeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   499
      unsafeEs = unsafeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   500
      swrappers = swrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   501
      uwrappers = uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   502
      unsafe_netpair = unsafe_netpair,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   503
      dup_netpair = dup_netpair,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   504
      extra_netpair = delete ([], [th]) extra_netpair}
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   505
  end;
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   506
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   507
fun del_unsafe_intro (r as (th, (th', swapped_th'), (dup_th', swapped_dup_th')))
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   508
  (CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   509
    safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   510
  CS
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   511
   {unsafe_netpair = delete ([th'], swapped_th') unsafe_netpair,
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   512
    dup_netpair = delete ([dup_th'], swapped_dup_th') dup_netpair,
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   513
    safeIs = safeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   514
    safeEs = safeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   515
    unsafeIs = Item_Net.remove r unsafeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   516
    unsafeEs = unsafeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   517
    swrappers = swrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   518
    uwrappers = uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   519
    safe0_netpair = safe0_netpair,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   520
    safep_netpair = safep_netpair,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   521
    extra_netpair = delete ([th], []) extra_netpair};
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   522
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   523
fun del_unsafe_elim (r as (th, (th', _), (dup_th', _)))
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   524
  (CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   525
    safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   526
  CS
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   527
   {unsafe_netpair = delete ([], [th']) unsafe_netpair,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   528
    dup_netpair = delete ([], [dup_th']) dup_netpair,
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   529
    safeIs = safeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   530
    safeEs = safeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   531
    unsafeIs = unsafeIs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   532
    unsafeEs = Item_Net.remove r unsafeEs,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   533
    swrappers = swrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   534
    uwrappers = uwrappers,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   535
    safe0_netpair = safe0_netpair,
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   536
    safep_netpair = safep_netpair,
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   537
    extra_netpair = delete ([], [th]) extra_netpair};
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   538
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   539
fun del f rules th cs =
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   540
  fold f (Item_Net.lookup rules (th, (th, []), (th, []))) cs;
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   541
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   542
in
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   543
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   544
fun delrule ctxt th (cs as CS {safeIs, safeEs, unsafeIs, unsafeEs, ...}) =
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   545
  let
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   546
    val th' = Tactic.make_elim th;
61055
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   547
    val r = (th, (th, []), (th, []));
6fc78876f9be store result of swapify, to avoid later access to implicit context;
wenzelm
parents: 61049
diff changeset
   548
    val r' = (th', (th', []), (th', []));
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   549
  in
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   550
    if Item_Net.member safeIs r orelse Item_Net.member safeEs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   551
      Item_Net.member unsafeIs r orelse Item_Net.member unsafeEs r orelse
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   552
      Item_Net.member safeEs r' orelse Item_Net.member unsafeEs r'
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   553
    then
60942
wenzelm
parents: 60817
diff changeset
   554
      cs
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   555
      |> del del_safe_intro safeIs th
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   556
      |> del del_safe_elim safeEs th
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   557
      |> del del_safe_elim safeEs th'
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   558
      |> del del_unsafe_intro unsafeIs th
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   559
      |> del del_unsafe_elim unsafeEs th
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   560
      |> del del_unsafe_elim unsafeEs th'
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   561
    else (warn_thm ctxt "Undeclared classical rule\n" th; cs)
9938
cb6a7572d0a1 delrule: handle dest rules as well;
wenzelm
parents: 9899
diff changeset
   562
  end;
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   563
60946
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   564
end;
46ec72073dc1 delete precisely the added rules;
wenzelm
parents: 60945
diff changeset
   565
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   566
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   567
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   568
(** claset data **)
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   569
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   570
(* wrappers *)
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   571
22674
1a610904bbca canonical merge operations
haftmann
parents: 22474
diff changeset
   572
fun map_swrappers f
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   573
  (CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   574
    safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   575
  CS {safeIs = safeIs, safeEs = safeEs, unsafeIs = unsafeIs, unsafeEs = unsafeEs,
4767
b9f3468c6ee2 introduced functions for updating the wrapper lists
oheimb
parents: 4765
diff changeset
   576
    swrappers = f swrappers, uwrappers = uwrappers,
b9f3468c6ee2 introduced functions for updating the wrapper lists
oheimb
parents: 4765
diff changeset
   577
    safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   578
    unsafe_netpair = unsafe_netpair, dup_netpair = dup_netpair, extra_netpair = extra_netpair};
4767
b9f3468c6ee2 introduced functions for updating the wrapper lists
oheimb
parents: 4765
diff changeset
   579
22674
1a610904bbca canonical merge operations
haftmann
parents: 22474
diff changeset
   580
fun map_uwrappers f
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   581
  (CS {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers,
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   582
    safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   583
  CS {safeIs = safeIs, safeEs = safeEs, unsafeIs = unsafeIs, unsafeEs = unsafeEs,
4767
b9f3468c6ee2 introduced functions for updating the wrapper lists
oheimb
parents: 4765
diff changeset
   584
    swrappers = swrappers, uwrappers = f uwrappers,
b9f3468c6ee2 introduced functions for updating the wrapper lists
oheimb
parents: 4765
diff changeset
   585
    safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   586
    unsafe_netpair = unsafe_netpair, dup_netpair = dup_netpair, extra_netpair = extra_netpair};
4767
b9f3468c6ee2 introduced functions for updating the wrapper lists
oheimb
parents: 4765
diff changeset
   587
22674
1a610904bbca canonical merge operations
haftmann
parents: 22474
diff changeset
   588
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   589
(* merge_cs *)
982
4fe0b642b7d5 Addition of wrappers for integration with the simplifier.
lcp
parents: 747
diff changeset
   590
42810
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   591
(*Merge works by adding all new rules of the 2nd claset into the 1st claset,
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   592
  in order to preserve priorities reliably.*)
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   593
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   594
fun merge_thms add thms1 thms2 =
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   595
  fold_rev (fn thm => if Item_Net.member thms1 thm then I else add thm) (Item_Net.content thms2);
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   596
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   597
fun merge_cs (cs as CS {safeIs, safeEs, unsafeIs, unsafeEs, ...},
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   598
    cs' as CS {safeIs = safeIs2, safeEs = safeEs2, unsafeIs = unsafeIs2, unsafeEs = unsafeEs2,
22674
1a610904bbca canonical merge operations
haftmann
parents: 22474
diff changeset
   599
      swrappers, uwrappers, ...}) =
24358
d75af3e90e82 tuned merge operations via pointer_eq;
wenzelm
parents: 24218
diff changeset
   600
  if pointer_eq (cs, cs') then cs
d75af3e90e82 tuned merge operations via pointer_eq;
wenzelm
parents: 24218
diff changeset
   601
  else
42810
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   602
    cs
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   603
    |> merge_thms (add_safe_intro NONE) safeIs safeIs2
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   604
    |> merge_thms (add_safe_elim NONE) safeEs safeEs2
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   605
    |> merge_thms (add_unsafe_intro NONE) unsafeIs unsafeIs2
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   606
    |> merge_thms (add_unsafe_elim NONE) unsafeEs unsafeEs2
42810
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   607
    |> map_swrappers (fn ws => AList.merge (op =) (K true) (ws, swrappers))
2425068fe13a slightly more efficient claset operations, using Item_Net to maintain rules in canonical order;
wenzelm
parents: 42807
diff changeset
   608
    |> map_uwrappers (fn ws => AList.merge (op =) (K true) (ws, uwrappers));
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   609
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   610
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   611
(* data *)
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   612
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   613
structure Claset = Generic_Data
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   614
(
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   615
  type T = claset;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   616
  val empty = empty_cs;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   617
  val extend = I;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   618
  val merge = merge_cs;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   619
);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   620
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   621
val claset_of = Claset.get o Context.Proof;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   622
val rep_claset_of = rep_cs o claset_of;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   623
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   624
val get_cs = Claset.get;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   625
val map_cs = Claset.map;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   626
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   627
fun map_theory_claset f thy =
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   628
  let
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   629
    val ctxt' = f (Proof_Context.init_global thy);
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   630
    val thy' = Proof_Context.theory_of ctxt';
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   631
  in Context.theory_map (Claset.map (K (claset_of ctxt'))) thy' end;
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   632
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   633
fun map_claset f = Context.proof_map (map_cs f);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   634
fun put_claset cs = map_claset (K cs);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   635
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   636
fun print_claset ctxt =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   637
  let
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   638
    val {safeIs, safeEs, unsafeIs, unsafeEs, swrappers, uwrappers, ...} = rep_claset_of ctxt;
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61056
diff changeset
   639
    val pretty_thms = map (Thm.pretty_thm_item ctxt o #1) o Item_Net.content;
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   640
  in
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   641
    [Pretty.big_list "safe introduction rules (intro!):" (pretty_thms safeIs),
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   642
      Pretty.big_list "introduction rules (intro):" (pretty_thms unsafeIs),
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   643
      Pretty.big_list "safe elimination rules (elim!):" (pretty_thms safeEs),
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   644
      Pretty.big_list "elimination rules (elim):" (pretty_thms unsafeEs),
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   645
      Pretty.strs ("safe wrappers:" :: map #1 swrappers),
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   646
      Pretty.strs ("unsafe wrappers:" :: map #1 uwrappers)]
56334
6b3739fee456 some shortcuts for chunks, which sometimes avoid bulky string output;
wenzelm
parents: 54742
diff changeset
   647
    |> Pretty.writeln_chunks
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   648
  end;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   649
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   650
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   651
(* old-style declarations *)
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   652
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   653
fun decl f (ctxt, ths) = map_claset (fold_rev (f ctxt) ths) ctxt;
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   654
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   655
val op addSIs = decl (addSI NONE);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   656
val op addSEs = decl (addSE NONE);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   657
val op addSDs = decl (addSD NONE);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   658
val op addIs = decl (addI NONE);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   659
val op addEs = decl (addE NONE);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   660
val op addDs = decl (addD NONE);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   661
val op delrules = decl delrule;
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   662
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   663
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   664
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   665
(*** Modifying the wrapper tacticals ***)
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   666
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   667
fun appSWrappers ctxt = fold (fn (_, w) => w ctxt) (#swrappers (rep_claset_of ctxt));
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   668
fun appWrappers ctxt = fold (fn (_, w) => w ctxt) (#uwrappers (rep_claset_of ctxt));
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   669
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   670
fun update_warn msg (p as (key : string, _)) xs =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   671
  (if AList.defined (op =) xs key then warning msg else (); AList.update (op =) p xs);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   672
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   673
fun delete_warn msg (key : string) xs =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   674
  if AList.defined (op =) xs key then AList.delete (op =) key xs
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   675
  else (warning msg; xs);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   676
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   677
(*Add/replace a safe wrapper*)
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   678
fun ctxt addSWrapper new_swrapper = ctxt |> map_claset
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   679
  (map_swrappers (update_warn ("Overwriting safe wrapper " ^ fst new_swrapper) new_swrapper));
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   680
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   681
(*Add/replace an unsafe wrapper*)
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   682
fun ctxt addWrapper new_uwrapper = ctxt |> map_claset
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   683
  (map_uwrappers (update_warn ("Overwriting unsafe wrapper " ^ fst new_uwrapper) new_uwrapper));
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   684
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   685
(*Remove a safe wrapper*)
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   686
fun ctxt delSWrapper name = ctxt |> map_claset
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   687
  (map_swrappers (delete_warn ("No such safe wrapper in claset: " ^ name) name));
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   688
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   689
(*Remove an unsafe wrapper*)
51703
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   690
fun ctxt delWrapper name = ctxt |> map_claset
f2e92fc0c8aa modifiers for classical wrappers operate on Proof.context instead of claset;
wenzelm
parents: 51688
diff changeset
   691
  (map_uwrappers (delete_warn ("No such unsafe wrapper in claset: " ^ name) name));
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   692
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   693
(* compose a safe tactic alternatively before/after safe_step_tac *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   694
fun ctxt addSbefore (name, tac1) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   695
  ctxt addSWrapper (name, fn ctxt => fn tac2 => tac1 ctxt ORELSE' tac2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   696
fun ctxt addSafter (name, tac2) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   697
  ctxt addSWrapper (name, fn ctxt => fn tac1 => tac1 ORELSE' tac2 ctxt);
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   698
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   699
(*compose a tactic alternatively before/after the step tactic *)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   700
fun ctxt addbefore (name, tac1) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   701
  ctxt addWrapper (name, fn ctxt => fn tac2 => tac1 ctxt APPEND' tac2);
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   702
fun ctxt addafter (name, tac2) =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51703
diff changeset
   703
  ctxt addWrapper (name, fn ctxt => fn tac1 => tac1 APPEND' tac2 ctxt);
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   704
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   705
fun ctxt addD2 (name, thm) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   706
  ctxt addafter (name, fn ctxt' => dresolve_tac ctxt' [thm] THEN' assume_tac ctxt');
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   707
fun ctxt addE2 (name, thm) =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   708
  ctxt addafter (name, fn ctxt' => eresolve_tac ctxt' [thm] THEN' assume_tac ctxt');
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   709
fun ctxt addSD2 (name, thm) =
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   710
  ctxt addSafter (name, fn ctxt' => dmatch_tac ctxt' [thm] THEN' eq_assume_tac);
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   711
fun ctxt addSE2 (name, thm) =
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   712
  ctxt addSafter (name, fn ctxt' => ematch_tac ctxt' [thm] THEN' eq_assume_tac);
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   713
1711
c06d01f75764 Provides merge_cs to support default clasets
paulson
parents: 1587
diff changeset
   714
982
4fe0b642b7d5 Addition of wrappers for integration with the simplifier.
lcp
parents: 747
diff changeset
   715
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   716
(**** Simple tactics for theorem proving ****)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   717
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   718
(*Attack subgoals using safe inferences -- matching, not resolution*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   719
fun safe_step_tac ctxt =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   720
  let val {safe0_netpair, safep_netpair, ...} = rep_claset_of ctxt in
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   721
    appSWrappers ctxt
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   722
      (FIRST'
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   723
       [eq_assume_tac,
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   724
        eq_mp_tac ctxt,
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   725
        bimatch_from_nets_tac ctxt safe0_netpair,
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51718
diff changeset
   726
        FIRST' (map (fn tac => tac ctxt) Data.hyp_subst_tacs),
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   727
        bimatch_from_nets_tac ctxt safep_netpair])
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   728
  end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   729
5757
0ad476dabbc6 corrected (and simplified) depth_tac
oheimb
parents: 5523
diff changeset
   730
(*Repeatedly attack a subgoal using safe inferences -- it's deterministic!*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   731
fun safe_steps_tac ctxt =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   732
  REPEAT_DETERM1 o (fn i => COND (has_fewer_prems i) no_tac (safe_step_tac ctxt i));
5757
0ad476dabbc6 corrected (and simplified) depth_tac
oheimb
parents: 5523
diff changeset
   733
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   734
(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   735
fun safe_tac ctxt = REPEAT_DETERM1 (FIRSTGOAL (safe_steps_tac ctxt));
747
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   736
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   737
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   738
(*** Clarify_tac: do safe steps without causing branching ***)
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   739
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   740
fun nsubgoalsP n (k, brl) = (subgoals_of_brl brl = n);
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   741
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   742
(*version of bimatch_from_nets_tac that only applies rules that
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   743
  create precisely n subgoals.*)
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   744
fun n_bimatch_from_nets_tac ctxt n =
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   745
  biresolution_from_nets_tac ctxt (order_list o filter (nsubgoalsP n)) true;
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   746
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   747
fun eq_contr_tac ctxt i = ematch_tac ctxt [Data.not_elim] i THEN eq_assume_tac i;
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   748
fun eq_assume_contr_tac ctxt = eq_assume_tac ORELSE' eq_contr_tac ctxt;
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   749
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   750
(*Two-way branching is allowed only if one of the branches immediately closes*)
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   751
fun bimatch2_tac ctxt netpair i =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   752
  n_bimatch_from_nets_tac ctxt 2 netpair i THEN
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   753
  (eq_assume_contr_tac ctxt i ORELSE eq_assume_contr_tac ctxt (i + 1));
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   754
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   755
(*Attack subgoals using safe inferences -- matching, not resolution*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   756
fun clarify_step_tac ctxt =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   757
  let val {safe0_netpair, safep_netpair, ...} = rep_claset_of ctxt in
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   758
    appSWrappers ctxt
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   759
     (FIRST'
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   760
       [eq_assume_contr_tac ctxt,
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   761
        bimatch_from_nets_tac ctxt safe0_netpair,
51798
ad3a241def73 uniform Proof.context for hyp_subst_tac;
wenzelm
parents: 51718
diff changeset
   762
        FIRST' (map (fn tac => tac ctxt) Data.hyp_subst_tacs),
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   763
        n_bimatch_from_nets_tac ctxt 1 safep_netpair,
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58950
diff changeset
   764
        bimatch2_tac ctxt safep_netpair])
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   765
  end;
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   766
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   767
fun clarify_tac ctxt = SELECT_GOAL (REPEAT_DETERM (clarify_step_tac ctxt 1));
3705
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   768
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   769
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   770
(*** Unsafe steps instantiate variables or lose information ***)
76f3b2803982 Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
paulson
parents: 3546
diff changeset
   771
4066
7b508ac609f7 Fixed comments
paulson
parents: 3727
diff changeset
   772
(*Backtracking is allowed among the various these unsafe ways of
7b508ac609f7 Fixed comments
paulson
parents: 3727
diff changeset
   773
  proving a subgoal.  *)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   774
fun inst0_step_tac ctxt =
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
   775
  assume_tac ctxt APPEND'
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58958
diff changeset
   776
  contr_tac ctxt APPEND'
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   777
  biresolve_from_nets_tac ctxt (#safe0_netpair (rep_claset_of ctxt));
747
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   778
4066
7b508ac609f7 Fixed comments
paulson
parents: 3727
diff changeset
   779
(*These unsafe steps could generate more subgoals.*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   780
fun instp_step_tac ctxt =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   781
  biresolve_from_nets_tac ctxt (#safep_netpair (rep_claset_of ctxt));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   782
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   783
(*These steps could instantiate variables and are therefore unsafe.*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   784
fun inst_step_tac ctxt = inst0_step_tac ctxt APPEND' instp_step_tac ctxt;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   785
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   786
fun unsafe_step_tac ctxt =
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   787
  biresolve_from_nets_tac ctxt (#unsafe_netpair (rep_claset_of ctxt));
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   788
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   789
(*Single step for the prover.  FAILS unless it makes progress. *)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   790
fun step_tac ctxt i =
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   791
  safe_tac ctxt ORELSE appWrappers ctxt (inst_step_tac ctxt ORELSE' unsafe_step_tac ctxt) i;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   792
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   793
(*Using a "safe" rule to instantiate variables is unsafe.  This tactic
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   794
  allows backtracking from "safe" rules to "unsafe" rules here.*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   795
fun slow_step_tac ctxt i =
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   796
  safe_tac ctxt ORELSE appWrappers ctxt (inst_step_tac ctxt APPEND' unsafe_step_tac ctxt) i;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   797
42791
36f787ae5f70 eliminated weight_ASTAR: int Unsynchronized.ref (astar_tac appears to be obsolete anyway);
wenzelm
parents: 42790
diff changeset
   798
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   799
(**** The following tactics all fail unless they solve one goal ****)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   800
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   801
(*Dumb but fast*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   802
fun fast_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   803
  Object_Logic.atomize_prems_tac ctxt THEN' SELECT_GOAL (DEPTH_SOLVE (step_tac ctxt 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   804
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   805
(*Slower but smarter than fast_tac*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   806
fun best_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   807
  Object_Logic.atomize_prems_tac ctxt THEN'
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   808
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, Data.sizef) (step_tac ctxt 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   809
9402
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 9184
diff changeset
   810
(*even a bit smarter than best_tac*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   811
fun first_best_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   812
  Object_Logic.atomize_prems_tac ctxt THEN'
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   813
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, Data.sizef) (FIRSTGOAL (step_tac ctxt)));
9402
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 9184
diff changeset
   814
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   815
fun slow_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   816
  Object_Logic.atomize_prems_tac ctxt THEN'
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   817
  SELECT_GOAL (DEPTH_SOLVE (slow_step_tac ctxt 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   818
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   819
fun slow_best_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   820
  Object_Logic.atomize_prems_tac ctxt THEN'
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   821
  SELECT_GOAL (BEST_FIRST (has_fewer_prems 1, Data.sizef) (slow_step_tac ctxt 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   822
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   823
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   824
(***ASTAR with weight weight_ASTAR, by Norbert Voelker*)
42791
36f787ae5f70 eliminated weight_ASTAR: int Unsynchronized.ref (astar_tac appears to be obsolete anyway);
wenzelm
parents: 42790
diff changeset
   825
36f787ae5f70 eliminated weight_ASTAR: int Unsynchronized.ref (astar_tac appears to be obsolete anyway);
wenzelm
parents: 42790
diff changeset
   826
val weight_ASTAR = 5;
1587
e7d8a4957bac Now provides astar versions (thanks to Norbert Voelker)
paulson
parents: 1524
diff changeset
   827
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   828
fun astar_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   829
  Object_Logic.atomize_prems_tac ctxt THEN'
10382
1fb807260ff1 atomize: all automated tactics that "solve" goals;
wenzelm
parents: 10309
diff changeset
   830
  SELECT_GOAL
52462
a241826ed003 actually use Data.sizef, not hardwired size_of_thm;
wenzelm
parents: 51798
diff changeset
   831
    (ASTAR (has_fewer_prems 1, fn lev => fn thm => Data.sizef thm + weight_ASTAR * lev)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   832
      (step_tac ctxt 1));
1587
e7d8a4957bac Now provides astar versions (thanks to Norbert Voelker)
paulson
parents: 1524
diff changeset
   833
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   834
fun slow_astar_tac ctxt =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   835
  Object_Logic.atomize_prems_tac ctxt THEN'
10382
1fb807260ff1 atomize: all automated tactics that "solve" goals;
wenzelm
parents: 10309
diff changeset
   836
  SELECT_GOAL
52462
a241826ed003 actually use Data.sizef, not hardwired size_of_thm;
wenzelm
parents: 51798
diff changeset
   837
    (ASTAR (has_fewer_prems 1, fn lev => fn thm => Data.sizef thm + weight_ASTAR * lev)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   838
      (slow_step_tac ctxt 1));
1587
e7d8a4957bac Now provides astar versions (thanks to Norbert Voelker)
paulson
parents: 1524
diff changeset
   839
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   840
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   841
(**** Complete tactic, loosely based upon LeanTaP.  This tactic is the outcome
747
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   842
  of much experimentation!  Changing APPEND to ORELSE below would prove
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   843
  easy theorems faster, but loses completeness -- and many of the harder
1800
3d9d2ef0cd3b Now implements delrules
paulson
parents: 1724
diff changeset
   844
  theorems such as 43. ****)
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   845
747
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   846
(*Non-deterministic!  Could always expand the first unsafe connective.
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   847
  That's hard to implement and did not perform better in experiments, due to
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   848
  greater search depth required.*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   849
fun dup_step_tac ctxt =
59164
ff40c53d1af9 proper context for "net" tactics;
wenzelm
parents: 59119
diff changeset
   850
  biresolve_from_nets_tac ctxt (#dup_netpair (rep_claset_of ctxt));
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   851
5523
dc8cdc192cd0 added addD2, addE2, addSD2, and addSE2
oheimb
parents: 5028
diff changeset
   852
(*Searching to depth m. A variant called nodup_depth_tac appears in clasimp.ML*)
5757
0ad476dabbc6 corrected (and simplified) depth_tac
oheimb
parents: 5523
diff changeset
   853
local
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   854
  fun slow_step_tac' ctxt = appWrappers ctxt (instp_step_tac ctxt APPEND' dup_step_tac ctxt);
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   855
in
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   856
  fun depth_tac ctxt m i state = SELECT_GOAL
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   857
    (safe_steps_tac ctxt 1 THEN_ELSE
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   858
      (DEPTH_SOLVE (depth_tac ctxt m 1),
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   859
        inst0_step_tac ctxt 1 APPEND COND (K (m = 0)) no_tac
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   860
          (slow_step_tac' ctxt 1 THEN DEPTH_SOLVE (depth_tac ctxt (m - 1) 1)))) i state;
5757
0ad476dabbc6 corrected (and simplified) depth_tac
oheimb
parents: 5523
diff changeset
   861
end;
747
bdc066781063 deepen_tac: modified due to outcome of experiments. Its
lcp
parents: 681
diff changeset
   862
10736
7f94cb4517fa recover_order for single step tules;
wenzelm
parents: 10394
diff changeset
   863
(*Search, with depth bound m.
2173
08c68550460b Added a comment
paulson
parents: 2066
diff changeset
   864
  This is the "entry point", which does safe inferences first.*)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   865
fun safe_depth_tac ctxt m = SUBGOAL (fn (prem, i) =>
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   866
  let
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   867
    val deti = (*No Vars in the goal?  No need to backtrack between goals.*)
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   868
      if exists_subterm (fn Var _ => true | _ => false) prem then DETERM else I;
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   869
  in
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   870
    SELECT_GOAL (TRY (safe_tac ctxt) THEN DEPTH_SOLVE (deti (depth_tac ctxt m 1))) i
42790
e07e56300faa misc tuning and simplification;
wenzelm
parents: 42439
diff changeset
   871
  end);
681
9b02474744ca Provers/classical: now takes theorem "classical" as argument, proves "swap"
lcp
parents: 469
diff changeset
   872
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   873
fun deepen_tac ctxt = DEEPEN (2, 10) (safe_depth_tac ctxt);
24021
491c68f40bc4 added get_cs/map_cs;
wenzelm
parents: 23594
diff changeset
   874
491c68f40bc4 added get_cs/map_cs;
wenzelm
parents: 23594
diff changeset
   875
5885
6c920d876098 tuned attribute names;
wenzelm
parents: 5841
diff changeset
   876
(* attributes *)
6c920d876098 tuned attribute names;
wenzelm
parents: 5841
diff changeset
   877
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   878
fun attrib f =
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   879
  Thm.declaration_attribute (fn th => fn context =>
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   880
    map_cs (f (Context.proof_of context) th) context);
5885
6c920d876098 tuned attribute names;
wenzelm
parents: 5841
diff changeset
   881
18691
a2dc15d9d6c8 attributes: optional weight;
wenzelm
parents: 18688
diff changeset
   882
val safe_elim = attrib o addSE;
a2dc15d9d6c8 attributes: optional weight;
wenzelm
parents: 18688
diff changeset
   883
val safe_intro = attrib o addSI;
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   884
val safe_dest = attrib o addSD;
60943
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   885
val unsafe_elim = attrib o addE;
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   886
val unsafe_intro = attrib o addI;
7cf1ea00a020 tuned signature;
wenzelm
parents: 60942
diff changeset
   887
val unsafe_dest = attrib o addD;
45375
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 42817
diff changeset
   888
7fe19930dfc9 more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
wenzelm
parents: 42817
diff changeset
   889
val rule_del =
60945
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   890
  Thm.declaration_attribute (fn th => fn context =>
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   891
    context
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   892
    |> map_cs (delrule (Context.proof_of context) th)
88aaecbaf61e clarified context;
wenzelm
parents: 60944
diff changeset
   893
    |> Thm.attribute_declaration Context_Rules.rule_del th);
5885
6c920d876098 tuned attribute names;
wenzelm
parents: 5841
diff changeset
   894
6c920d876098 tuned attribute names;
wenzelm
parents: 5841
diff changeset
   895
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   896
5885
6c920d876098 tuned attribute names;
wenzelm
parents: 5841
diff changeset
   897
(** concrete syntax of attributes **)
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   898
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   899
val introN = "intro";
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   900
val elimN = "elim";
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   901
val destN = "dest";
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   902
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   903
val _ =
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   904
  Theory.setup
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   905
   (Attrib.setup \<^binding>\<open>swapped\<close> (Scan.succeed swapped)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   906
      "classical swap of introduction rule" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   907
    Attrib.setup \<^binding>\<open>dest\<close> (Context_Rules.add safe_dest unsafe_dest Context_Rules.dest_query)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   908
      "declaration of Classical destruction rule" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   909
    Attrib.setup \<^binding>\<open>elim\<close> (Context_Rules.add safe_elim unsafe_elim Context_Rules.elim_query)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   910
      "declaration of Classical elimination rule" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   911
    Attrib.setup \<^binding>\<open>intro\<close> (Context_Rules.add safe_intro unsafe_intro Context_Rules.intro_query)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   912
      "declaration of Classical introduction rule" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   913
    Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   914
      "remove declaration of intro/elim/dest rule");
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   915
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   916
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   917
7230
4ca0d7839ff1 renamed 'single' to 'some_rule';
wenzelm
parents: 7154
diff changeset
   918
(** proof methods **)
4ca0d7839ff1 renamed 'single' to 'some_rule';
wenzelm
parents: 7154
diff changeset
   919
4ca0d7839ff1 renamed 'single' to 'some_rule';
wenzelm
parents: 7154
diff changeset
   920
local
4ca0d7839ff1 renamed 'single' to 'some_rule';
wenzelm
parents: 7154
diff changeset
   921
30609
983e8b6e4e69 Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
wenzelm
parents: 30558
diff changeset
   922
fun some_rule_tac ctxt facts = SUBGOAL (fn (goal, i) =>
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   923
  let
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 60946
diff changeset
   924
    val [rules1, rules2, rules4] = Context_Rules.find_rules ctxt false facts goal;
60942
wenzelm
parents: 60817
diff changeset
   925
    val {extra_netpair, ...} = rep_claset_of ctxt;
61049
0d401f874942 trim context for persistent storage;
wenzelm
parents: 60946
diff changeset
   926
    val rules3 = Context_Rules.find_rules_netpair ctxt true facts goal extra_netpair;
12376
480303e3fa08 simplified (and clarified) integration with Pure/ContextRules;
wenzelm
parents: 12362
diff changeset
   927
    val rules = rules1 @ rules2 @ rules3 @ rules4;
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 58893
diff changeset
   928
    val ruleq = Drule.multi_resolves (SOME ctxt) facts rules;
52732
b4da1f2ec73f standardized aliases;
wenzelm
parents: 52699
diff changeset
   929
    val _ = Method.trace ctxt rules;
12376
480303e3fa08 simplified (and clarified) integration with Pure/ContextRules;
wenzelm
parents: 12362
diff changeset
   930
  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59164
diff changeset
   931
    fn st => Seq.maps (fn rule => resolve_tac ctxt [rule] i st) ruleq
18834
7e94af77cfce default rule step: norm_hhf_tac;
wenzelm
parents: 18728
diff changeset
   932
  end)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   933
  THEN_ALL_NEW Goal.norm_hhf_tac ctxt;
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   934
30609
983e8b6e4e69 Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
wenzelm
parents: 30558
diff changeset
   935
in
7281
2f454e1fd372 renamed 'some_rule' to 'rule';
wenzelm
parents: 7272
diff changeset
   936
30609
983e8b6e4e69 Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
wenzelm
parents: 30558
diff changeset
   937
fun rule_tac ctxt [] facts = some_rule_tac ctxt facts
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52732
diff changeset
   938
  | rule_tac ctxt rules facts = Method.rule_tac ctxt rules facts;
30609
983e8b6e4e69 Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
wenzelm
parents: 30558
diff changeset
   939
60619
7512716f03db no arguments for "standard" (or old "default") methods;
wenzelm
parents: 60618
diff changeset
   940
fun standard_tac ctxt facts =
7512716f03db no arguments for "standard" (or old "default") methods;
wenzelm
parents: 60618
diff changeset
   941
  HEADGOAL (some_rule_tac ctxt facts) ORELSE
60618
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 60097
diff changeset
   942
  Class.standard_intro_classes_tac ctxt facts;
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 60097
diff changeset
   943
7230
4ca0d7839ff1 renamed 'single' to 'some_rule';
wenzelm
parents: 7154
diff changeset
   944
end;
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   945
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   946
6502
bc30e13b36a8 improved 'single' method;
wenzelm
parents: 6492
diff changeset
   947
(* automatic methods *)
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   948
5927
991483daa1a4 expoer cla_method('), cla_modifiers;
wenzelm
parents: 5885
diff changeset
   949
val cla_modifiers =
64556
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   950
 [Args.$$$ destN -- Args.bang_colon >> K (Method.modifier (safe_dest NONE) \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   951
  Args.$$$ destN -- Args.colon >> K (Method.modifier (unsafe_dest NONE) \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   952
  Args.$$$ elimN -- Args.bang_colon >> K (Method.modifier (safe_elim NONE) \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   953
  Args.$$$ elimN -- Args.colon >> K (Method.modifier (unsafe_elim NONE) \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   954
  Args.$$$ introN -- Args.bang_colon >> K (Method.modifier (safe_intro NONE) \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   955
  Args.$$$ introN -- Args.colon >> K (Method.modifier (unsafe_intro NONE) \<^here>),
851ae0e7b09c more symbols;
wenzelm
parents: 62939
diff changeset
   956
  Args.del -- Args.colon >> K (Method.modifier rule_del \<^here>)];
5927
991483daa1a4 expoer cla_method('), cla_modifiers;
wenzelm
parents: 5885
diff changeset
   957
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   958
fun cla_method tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD o tac);
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42792
diff changeset
   959
fun cla_method' tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD' o tac);
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   960
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   961
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   962
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   963
(** method setup **)
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   964
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   965
val _ =
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   966
  Theory.setup
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   967
   (Method.setup \<^binding>\<open>standard\<close> (Scan.succeed (METHOD o standard_tac))
60618
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 60097
diff changeset
   968
      "standard proof step: classical intro/elim rule or class introduction" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   969
    Method.setup \<^binding>\<open>rule\<close>
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   970
      (Attrib.thms >> (fn rules => fn ctxt => METHOD (HEADGOAL o rule_tac ctxt rules)))
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   971
      "apply some intro/elim rule (potentially classical)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   972
    Method.setup \<^binding>\<open>contradiction\<close>
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   973
      (Scan.succeed (fn ctxt => Method.rule ctxt [Data.not_elim, Drule.rotate_prems 1 Data.not_elim]))
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   974
      "proof by contradiction" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   975
    Method.setup \<^binding>\<open>clarify\<close> (cla_method' (CHANGED_PROP oo clarify_tac))
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   976
      "repeatedly apply safe steps" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   977
    Method.setup \<^binding>\<open>fast\<close> (cla_method' fast_tac) "classical prover (depth-first)" #>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   978
    Method.setup \<^binding>\<open>slow\<close> (cla_method' slow_tac) "classical prover (slow depth-first)" #>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   979
    Method.setup \<^binding>\<open>best\<close> (cla_method' best_tac) "classical prover (best-first)" #>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   980
    Method.setup \<^binding>\<open>deepen\<close>
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   981
      (Scan.lift (Scan.optional Parse.nat 4) --| Method.sections cla_modifiers
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   982
        >> (fn n => fn ctxt => SIMPLE_METHOD' (deepen_tac ctxt n)))
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   983
      "classical prover (iterative deepening)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   984
    Method.setup \<^binding>\<open>safe\<close> (cla_method (CHANGED_PROP o safe_tac))
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   985
      "classical prover (apply safe rules)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   986
    Method.setup \<^binding>\<open>safe_step\<close> (cla_method' safe_step_tac)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   987
      "single classical step (safe rules)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   988
    Method.setup \<^binding>\<open>inst_step\<close> (cla_method' inst_step_tac)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   989
      "single classical step (safe rules, allow instantiations)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   990
    Method.setup \<^binding>\<open>step\<close> (cla_method' step_tac)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   991
      "single classical step (safe and unsafe rules)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   992
    Method.setup \<^binding>\<open>slow_step\<close> (cla_method' slow_step_tac)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   993
      "single classical step (safe and unsafe rules, allow backtracking)" #>
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
   994
    Method.setup \<^binding>\<open>clarify_step\<close> (cla_method' clarify_step_tac)
58826
2ed2eaabe3df modernized setup;
wenzelm
parents: 58048
diff changeset
   995
      "single classical step (safe rules, without splitting)");
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
   996
8667
4230d17073ea print_simpset / print_claset command;
wenzelm
parents: 8519
diff changeset
   997
4230d17073ea print_simpset / print_claset command;
wenzelm
parents: 8519
diff changeset
   998
(** outer syntax **)
4230d17073ea print_simpset / print_claset command;
wenzelm
parents: 8519
diff changeset
   999
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24358
diff changeset
  1000
val _ =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67649
diff changeset
  1001
  Outer_Syntax.command \<^command_keyword>\<open>print_claset\<close> "print context of Classical Reasoner"
60097
d20ca79d50e4 discontinued pointless warnings: commands are only defined inside a theory context;
wenzelm
parents: 59970
diff changeset
  1002
    (Scan.succeed (Toplevel.keep (print_claset o Toplevel.context_of)));
8667
4230d17073ea print_simpset / print_claset command;
wenzelm
parents: 8519
diff changeset
  1003
5841
d05df8752a8a local claset theory data;
wenzelm
parents: 5757
diff changeset
  1004
end;