src/Pure/Isar/calculation.ML
author wenzelm
Thu, 06 Dec 2001 00:42:00 +0100
changeset 12402 cef751fff6b0
parent 12379 c74d160ff0c5
child 12805 3be853cf19cf
permissions -rw-r--r--
clarified sym_del;
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
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    17
  val sym_add_global: theory attribute
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    18
  val sym_del_global: theory attribute
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    19
  val sym_add_local: Proof.context attribute
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    20
  val sym_del_local: Proof.context attribute
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    21
  val symmetric_global: theory attribute
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    22
  val symmetric_local: Proof.context attribute
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    23
  val also: thm list option -> (Proof.context -> thm list -> unit)
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    24
    -> Proof.state -> Proof.state Seq.seq
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    25
  val finally: thm list option -> (Proof.context -> thm list -> unit)
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    26
    -> Proof.state -> Proof.state Seq.seq
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    27
  val moreover: (Proof.context -> thm list -> unit) -> Proof.state -> Proof.state
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    28
  val ultimately: (Proof.context -> thm list -> unit) -> Proof.state -> Proof.state
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    29
  val setup: (theory -> theory) list
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    30
end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    31
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    32
structure Calculation: CALCULATION =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    33
struct
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    34
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    35
(** global and local calculation data **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    36
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
    37
(* theory data kind 'Isar/calculation' *)
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    38
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    39
fun print_rules prt x (trans, sym) =
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    40
  [Pretty.big_list "transitivity rules:" (map (prt x) (NetRules.rules trans)),
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    41
    Pretty.big_list "symmetry rules:" (map (prt x) sym)]
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    42
  |> Pretty.chunks |> Pretty.writeln;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    43
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    44
structure GlobalCalculationArgs =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    45
struct
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    46
  val name = "Isar/calculation";
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    47
  type T = thm NetRules.T * thm list
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    48
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    49
  val empty = (NetRules.elim, []);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    50
  val copy = I;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    51
  val prep_ext = I;
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    52
  fun merge ((trans1, sym1), (trans2, sym2)) =
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    53
    (NetRules.merge (trans1, trans2), Drule.merge_rules (sym1, sym2));
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    54
  val print = print_rules Display.pretty_thm_sg;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    55
end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    56
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    57
structure GlobalCalculation = TheoryDataFun(GlobalCalculationArgs);
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    58
val print_global_rules = GlobalCalculation.print;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    59
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    60
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    61
(* proof data kind 'Isar/calculation' *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    62
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    63
structure LocalCalculationArgs =
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    64
struct
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    65
  val name = "Isar/calculation";
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    66
  type T = (thm NetRules.T * thm list) * (thm list * int) option;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    67
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    68
  fun init thy = (GlobalCalculation.get thy, None);
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
    69
  fun print ctxt (rs, _) = print_rules ProofContext.pretty_thm ctxt rs;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    70
end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    71
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    72
structure LocalCalculation = ProofDataFun(LocalCalculationArgs);
6787
wenzelm
parents: 6782
diff changeset
    73
val get_local_rules = #1 o LocalCalculation.get_st;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    74
val print_local_rules = LocalCalculation.print;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    75
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    76
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    77
(* access calculation *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    78
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    79
fun get_calculation state =
6787
wenzelm
parents: 6782
diff changeset
    80
  (case #2 (LocalCalculation.get_st state) of
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    81
    None => None
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    82
  | Some (thms, lev) => if lev = Proof.level state then Some thms else None);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    83
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    84
fun put_calculation thms state =
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
    85
  LocalCalculation.put_st (get_local_rules state, Some (thms, Proof.level state)) state;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    86
6787
wenzelm
parents: 6782
diff changeset
    87
fun reset_calculation state =
wenzelm
parents: 6782
diff changeset
    88
  LocalCalculation.put_st (get_local_rules state, None) state;
wenzelm
parents: 6782
diff changeset
    89
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    90
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    91
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    92
(** attributes **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    93
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    94
(* add/del rules *)
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    95
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    96
fun global_att f (x, thm) = (GlobalCalculation.map (f thm) x, thm);
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    97
fun local_att f (x, thm) = (LocalCalculation.map (apfst (f thm)) x, thm);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
    98
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
    99
val trans_add_global = global_att (apfst o NetRules.insert);
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   100
val trans_del_global = global_att (apfst o NetRules.delete);
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   101
val trans_add_local = local_att (apfst o NetRules.insert);
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   102
val trans_del_local = local_att (apfst o NetRules.delete);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   103
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   104
val sym_add_global = global_att (apsnd o Drule.add_rule) o ContextRules.elim_query_global None;
12402
cef751fff6b0 clarified sym_del;
wenzelm
parents: 12379
diff changeset
   105
val sym_del_global = global_att (apsnd o Drule.del_rule) o ContextRules.rule_del_global;
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   106
val sym_add_local = local_att (apsnd o Drule.add_rule) o ContextRules.elim_query_local None;
12402
cef751fff6b0 clarified sym_del;
wenzelm
parents: 12379
diff changeset
   107
val sym_del_local = local_att (apsnd o Drule.del_rule) o ContextRules.rule_del_local;
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   108
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   109
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   110
(* symmetry *)
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   111
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   112
fun gen_symmetric get_sym = Drule.rule_attribute (fn x => fn th =>
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   113
  (case Seq.chop (2, Method.multi_resolves [th] (get_sym x)) of
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   114
    ([th'], _) => th'
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   115
  | ([], _) => raise THM ("symmetric: no unifiers", 1, [th])
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   116
  | _ => raise THM ("symmetric: multiple unifiers", 1, [th])));
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   117
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   118
val symmetric_global = gen_symmetric (#2 o GlobalCalculation.get);
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   119
val symmetric_local = gen_symmetric (#2 o #1 o LocalCalculation.get);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   120
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   121
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   122
(* concrete syntax *)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   123
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   124
val trans_attr =
8634
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8588
diff changeset
   125
 (Attrib.add_del_args trans_add_global trans_del_global,
3f34637cb9c0 use Attrib.add_del_args;
wenzelm
parents: 8588
diff changeset
   126
  Attrib.add_del_args trans_add_local trans_del_local);
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   127
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   128
val sym_attr =
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   129
 (Attrib.add_del_args sym_add_global sym_del_global,
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   130
  Attrib.add_del_args sym_add_local sym_del_local);
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   131
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   132
6787
wenzelm
parents: 6782
diff changeset
   133
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   134
(** proof commands **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   135
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   136
(* maintain calculation register *)
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   137
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   138
val calculationN = "calculation";
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   139
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   140
fun maintain_calculation false calc state =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   141
      state
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   142
      |> put_calculation calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   143
      |> Proof.simple_have_thms calculationN calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   144
      |> Proof.reset_facts
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   145
  | maintain_calculation true calc state =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   146
      state
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   147
      |> reset_calculation
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   148
      |> Proof.reset_thms calculationN
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   149
      |> Proof.simple_have_thms "" calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   150
      |> Proof.chain;
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   151
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   152
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   153
(* 'also' and 'finally' *)
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   154
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   155
fun err_if state b msg = if b then raise Proof.STATE (msg, state) else ();
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   156
6877
3d5e5e6f9e20 also, finally: opt_rules;
wenzelm
parents: 6787
diff changeset
   157
fun calculate final opt_rules print state =
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   158
  let
9322
b5bd2709a2c2 eq_prop: strip_assums_concl;
wenzelm
parents: 9217
diff changeset
   159
    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
   160
    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
   161
    fun projection ths th = Library.exists (Library.curry eq_prop th) ths;
8300
4c3f83414de3 use NetRules;
wenzelm
parents: 8150
diff changeset
   162
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   163
    fun combine ths =
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   164
      (case opt_rules of Some rules => rules
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   165
      | None =>
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   166
          (case ths of [] => NetRules.rules (#1 (get_local_rules state))
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   167
          | th :: _ => NetRules.retrieve (#1 (get_local_rules state)) (strip_assums_concl th)))
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   168
      |> Seq.of_list |> Seq.map (Method.multi_resolve ths) |> Seq.flat
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   169
      |> Seq.filter (not o projection ths);
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
   170
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   171
    val facts = Proof.the_facts (Proof.assert_forward state);
6903
682f8a9ec75d improved errors;
wenzelm
parents: 6877
diff changeset
   172
    val (initial, calculations) =
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   173
      (case get_calculation state of
7414
9bc7797d1249 calculation: thm list;
wenzelm
parents: 7197
diff changeset
   174
        None => (true, Seq.single facts)
11097
c1be9f2dff4c more robust selection of calculational rules;
wenzelm
parents: 10478
diff changeset
   175
      | Some calc => (false, Seq.map single (combine (calc @ facts))));
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   176
  in
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   177
    err_if state (initial andalso final) "No calculation yet";
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   178
    err_if state (initial andalso is_some opt_rules) "Initial calculation -- no rules to be given";
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
   179
    calculations |> Seq.map (fn calc => (print (Proof.context_of state) calc;
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
   180
        state |> maintain_calculation final calc))
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   181
  end;
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   182
6782
adf099e851ed fixed "...": dest_arg;
wenzelm
parents: 6781
diff changeset
   183
fun also print = calculate false print;
adf099e851ed fixed "...": dest_arg;
wenzelm
parents: 6781
diff changeset
   184
fun finally print = calculate true print;
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   185
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   186
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   187
(* 'moreover' and 'ultimately' *)
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   188
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   189
fun collect final print state =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   190
  let
10478
97247fd6f1f8 Proof.assert_forward;
wenzelm
parents: 10008
diff changeset
   191
    val facts = Proof.the_facts (Proof.assert_forward state);
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   192
    val (initial, thms) =
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   193
      (case get_calculation state of
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   194
        None => (true, [])
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   195
      | Some thms => (false, thms));
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   196
    val calc = thms @ facts;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   197
  in
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   198
    err_if state (initial andalso final) "No calculation yet";
12055
a9c44895cc8c pretty/print functions with context;
wenzelm
parents: 12021
diff changeset
   199
    print (Proof.context_of state) calc;
8588
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   200
    state |> maintain_calculation final calc
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   201
  end;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   202
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   203
fun moreover print = collect false print;
b7c3f264f8ac added 'ultimately';
wenzelm
parents: 8562
diff changeset
   204
fun ultimately print = collect true print;
8562
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   205
ce0e2b8e8844 added 'moreover' command;
wenzelm
parents: 8461
diff changeset
   206
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   207
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   208
(** theory setup **)
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   209
12379
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   210
val setup =
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   211
 [GlobalCalculation.init, LocalCalculation.init,
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   212
  Attrib.add_attributes
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   213
   [("trans", trans_attr, "declaration of transitivity rule"),
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   214
    ("sym", sym_attr, "declaration of symmetry rule"),
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   215
    ("symmetric", (Attrib.no_args symmetric_global, Attrib.no_args symmetric_local),
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   216
      "resolution with symmetry rule")],
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   217
  #1 o PureThy.add_thms
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   218
   [(("", transitive_thm), [trans_add_global]),
c74d160ff0c5 added 'sym' and 'symmetric' atts;
wenzelm
parents: 12311
diff changeset
   219
    (("", symmetric_thm), [sym_add_global])]];
6778
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   220
2f66eea8a025 Support for calculational proofs.
wenzelm
parents:
diff changeset
   221
end;