src/Pure/Isar/calculation.ML
author wenzelm
Mon, 15 Oct 2001 20:36:48 +0200
changeset 11784 b66b198ee29a
parent 11097 c1be9f2dff4c
child 12021 8809efda06d3
permissions -rw-r--r--
tuned NetRules;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/calculation.ML
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
8807
wenzelm
parents: 8649
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     5
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     6
Support for calculational proofs.
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     7
*)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     8
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
     9
signature CALCULATION =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    10
sig
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    11
  val print_global_rules: theory -> unit
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    12
  val print_local_rules: Proof.context -> unit
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    13
  val trans_add_global: theory attribute
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    14
  val trans_del_global: theory attribute
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    15
  val trans_add_local: Proof.context attribute
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    16
  val trans_del_local: Proof.context attribute
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    17
  val also: thm list option -> (thm list -> unit) -> Proof.state -> Proof.state Seq.seq
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    18
  val finally: thm list option -> (thm list -> unit) -> Proof.state -> Proof.state Seq.seq
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
    19
  val moreover: (thm list -> unit) -> Proof.state -> Proof.state
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
    20
  val ultimately: (thm list -> unit) -> Proof.state -> Proof.state
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    21
  val setup: (theory -> theory) list
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    22
end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    23
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    24
structure Calculation: CALCULATION =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    25
struct
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    26
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    27
(** global and local calculation data **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    28
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    29
(* theory data kind 'Isar/calculation' *)
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    30
10008
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 9900
diff changeset
    31
fun print_rules sg rs = Pretty.writeln (Pretty.big_list "transitivity rules:"
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 9900
diff changeset
    32
  (map (Display.pretty_thm_sg sg) (NetRules.rules rs)));
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    33
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    34
structure GlobalCalculationArgs =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    35
struct
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    36
  val name = "Isar/calculation";
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    37
  type T = thm NetRules.T
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    38
11784
b66b198ee29a tuned NetRules;
wenzelm
parents: 11097
diff changeset
    39
  val empty = NetRules.elim;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    40
  val copy = I;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    41
  val prep_ext = I;
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    42
  val merge = NetRules.merge;
10008
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 9900
diff changeset
    43
  val print = print_rules;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    44
end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    45
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    46
structure GlobalCalculation = TheoryDataFun(GlobalCalculationArgs);
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    47
val print_global_rules = GlobalCalculation.print;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    48
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    49
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    50
(* proof data kind 'Isar/calculation' *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    51
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    52
structure LocalCalculationArgs =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    53
struct
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    54
  val name = "Isar/calculation";
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    55
  type T = thm NetRules.T * (thm list * int) option;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    56
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    57
  fun init thy = (GlobalCalculation.get thy, None);
10008
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 9900
diff changeset
    58
  fun print ctxt (rs, _) = print_rules (ProofContext.sign_of ctxt) rs;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    59
end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    60
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    61
structure LocalCalculation = ProofDataFun(LocalCalculationArgs);
6787
wenzelm
parents: 6782
diff changeset
    62
val get_local_rules = #1 o LocalCalculation.get_st;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    63
val print_local_rules = LocalCalculation.print;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    64
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    65
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    66
(* access calculation *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    67
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    68
fun get_calculation state =
6787
wenzelm
parents: 6782
diff changeset
    69
  (case #2 (LocalCalculation.get_st state) of
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    70
    None => None
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    71
  | Some (thms, lev) => if lev = Proof.level state then Some thms else None);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    72
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    73
fun put_calculation thms state =
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    74
  LocalCalculation.put_st (get_local_rules state, Some (thms, Proof.level state)) state;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    75
6787
wenzelm
parents: 6782
diff changeset
    76
fun reset_calculation state =
wenzelm
parents: 6782
diff changeset
    77
  LocalCalculation.put_st (get_local_rules state, None) state;
wenzelm
parents: 6782
diff changeset
    78
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    79
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    80
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    81
(** attributes **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    82
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    83
(* trans add/del *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    84
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    85
fun mk_att f g (x, thm) = (f (g thm) x, thm);
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    86
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    87
val trans_add_global = mk_att GlobalCalculation.map NetRules.insert;
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    88
val trans_del_global = mk_att GlobalCalculation.map NetRules.delete;
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    89
val trans_add_local = mk_att LocalCalculation.map (Library.apfst o NetRules.insert);
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    90
val trans_del_local = mk_att LocalCalculation.map (Library.apfst o NetRules.delete);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    91
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    92
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    93
(* concrete syntax *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    94
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    95
val trans_attr =
8634
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8588
diff changeset
    96
 (Attrib.add_del_args trans_add_global trans_del_global,
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8588
diff changeset
    97
  Attrib.add_del_args trans_add_local trans_del_local);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    98
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    99
6787
wenzelm
parents: 6782
diff changeset
   100
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   101
(** proof commands **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   102
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   103
(* maintain calculation register *)
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   104
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   105
val calculationN = "calculation";
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   106
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   107
fun maintain_calculation false calc state =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   108
      state
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   109
      |> put_calculation calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   110
      |> Proof.simple_have_thms calculationN calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   111
      |> Proof.reset_facts
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   112
  | maintain_calculation true calc state =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   113
      state
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   114
      |> reset_calculation
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   115
      |> Proof.reset_thms calculationN
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   116
      |> Proof.simple_have_thms "" calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   117
      |> Proof.chain;
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   118
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   119
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   120
(* 'also' and 'finally' *)
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   121
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   122
fun err_if state b msg = if b then raise Proof.STATE (msg, state) else ();
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   123
6877
3d5e5e6f9e20 also, finally: opt_rules;
wenzelm
parents: 6787
diff changeset
   124
fun calculate final opt_rules print state =
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   125
  let
9322
b5bd2709a2c2 eq_prop: strip_assums_concl;
wenzelm
parents: 9217
diff changeset
   126
    val strip_assums_concl = Logic.strip_assums_concl o #prop o Thm.rep_thm;
b5bd2709a2c2 eq_prop: strip_assums_concl;
wenzelm
parents: 9217
diff changeset
   127
    val eq_prop = op aconv o pairself (Pattern.eta_contract o strip_assums_concl);
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   128
    fun projection ths th = Library.exists (Library.curry eq_prop th) ths;
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
   129
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   130
    fun combine ths =
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   131
      (case opt_rules of Some rules => rules
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   132
      | None =>
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   133
          (case ths of [] => NetRules.rules (get_local_rules state)
11784
b66b198ee29a tuned NetRules;
wenzelm
parents: 11097
diff changeset
   134
          | th :: _ => NetRules.retrieve (get_local_rules state) (strip_assums_concl th)))
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   135
      |> Seq.of_list |> Seq.map (Method.multi_resolve ths) |> Seq.flat
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   136
      |> Seq.filter (not o projection ths);
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
   137
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   138
    val facts = Proof.the_facts (Proof.assert_forward state);
6903
682f8a9ec75d improved errors;
wenzelm
parents: 6877
diff changeset
   139
    val (initial, calculations) =
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   140
      (case get_calculation state of
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
   141
        None => (true, Seq.single facts)
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   142
      | Some calc => (false, Seq.map single (combine (calc @ facts))));
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   143
  in
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   144
    err_if state (initial andalso final) "No calculation yet";
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   145
    err_if state (initial andalso is_some opt_rules) "Initial calculation -- no rules to be given";
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   146
    calculations |> Seq.map (fn calc => (print calc; state |> maintain_calculation final calc))
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   147
  end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   148
6782
adf099e851ed fixed "...": dest_arg;
wenzelm
parents: 6781
diff changeset
   149
fun also print = calculate false print;
adf099e851ed fixed "...": dest_arg;
wenzelm
parents: 6781
diff changeset
   150
fun finally print = calculate true print;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   151
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   152
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   153
(* 'moreover' and 'ultimately' *)
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   154
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   155
fun collect final print state =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   156
  let
10478
97247fd6f1f8 Proof.assert_forward;
wenzelm
parents: 10008
diff changeset
   157
    val facts = Proof.the_facts (Proof.assert_forward state);
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   158
    val (initial, thms) =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   159
      (case get_calculation state of
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   160
        None => (true, [])
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   161
      | Some thms => (false, thms));
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   162
    val calc = thms @ facts;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   163
  in
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   164
    err_if state (initial andalso final) "No calculation yet";
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   165
    print calc;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   166
    state |> maintain_calculation final calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   167
  end;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   168
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   169
fun moreover print = collect false print;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   170
fun ultimately print = collect true print;
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   171
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   172
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   173
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   174
(** theory setup **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   175
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   176
val setup = [GlobalCalculation.init, LocalCalculation.init,
9900
8035a13c61a0 tuned msgs;
wenzelm
parents: 9322
diff changeset
   177
  Attrib.add_attributes [("trans", trans_attr, "declaration of transitivity rule")]];
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   178
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   179
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   180
end;