src/Provers/Arith/extract_common_term.ML
author wenzelm
Tue, 04 Jun 2019 20:49:33 +0200
changeset 70326 aa7c49651f4e
parent 70315 2f9623aa1eeb
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     1
(*  Title:      Provers/Arith/extract_common_term.ML
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     3
    Copyright   2000  University of Cambridge
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     4
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     5
Extract common terms in balanced expressions:
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     6
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     7
     i + u + j ~~ i' + u + j'  ==  u + (i + j) ~~ u + (i' + j')
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     8
     i + u     ~~ u            ==  u + i       ~~ u + 0
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     9
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    10
where ~~ is an appropriate balancing operation (e.g. =, <=, <, -) and 0 is a
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    11
suitable identity for +.
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    12
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    13
This massaged formula is then simplified in a user-specified way.
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    14
*)
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    15
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    16
signature EXTRACT_COMMON_TERM_DATA =
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    17
sig
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    18
  (*abstract syntax*)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 13484
diff changeset
    19
  val mk_sum: typ -> term list -> term
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    20
  val dest_sum: term -> term list
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    21
  val mk_bal: term * term -> term
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    22
  val dest_bal: term -> term * term
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    23
  val find_first: term -> term list -> term list
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    24
  (*proof tools*)
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 43597
diff changeset
    25
  val mk_eq: term * term -> term
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45306
diff changeset
    26
  val norm_tac: Proof.context -> tactic                (*proves the result*)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45306
diff changeset
    27
  val simplify_meta_eq: Proof.context -> thm -> thm -> thm (*simplifies the result*)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45306
diff changeset
    28
  val simp_conv: Proof.context -> term -> thm option  (*proves simp thm*)
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    29
end;
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    30
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    31
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    32
functor ExtractCommonTermFun(Data: EXTRACT_COMMON_TERM_DATA):
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    33
  sig
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45306
diff changeset
    34
  val proc: Proof.context -> term -> thm option
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    35
  end
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    36
=
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    37
struct
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    38
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    39
(*a left-to-right scan of terms1, seeking a term u that is also in terms2*)
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    40
fun find_common (terms1,terms2) =
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17223
diff changeset
    41
  let val tab2 = fold (Termtab.update o rpair ()) terms2 Termtab.empty
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    42
      fun seek [] = raise TERM("find_common", [])
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    43
        | seek (u::terms) =
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    44
              if Termtab.defined tab2 u then u
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    45
              else seek terms
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    46
  in seek terms1 end;
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    47
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    48
(*the simplification procedure*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45306
diff changeset
    49
fun proc ctxt t =
15027
d23887300b96 adapted type of simprocs;
wenzelm
parents: 14387
diff changeset
    50
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45306
diff changeset
    51
    val prems = Simplifier.prems_of ctxt;
70326
wenzelm
parents: 70315
diff changeset
    52
    val (t', ctxt') = yield_singleton (Variable.import_terms true) t ctxt
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    53
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    54
    val (t1,t2) = Data.dest_bal t'
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    55
    val terms1 = Data.dest_sum t1
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    56
    and terms2 = Data.dest_sum t2
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    57
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    58
    val u = find_common (terms1,terms2)
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 20114
diff changeset
    59
    val simp_th =
70315
2f9623aa1eeb proper context;
wenzelm
parents: 51717
diff changeset
    60
          case Data.simp_conv ctxt' u of NONE => raise TERM("no simp", [])
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 20114
diff changeset
    61
          | SOME th => th
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    62
    val terms1' = Data.find_first u terms1
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    63
    and terms2' = Data.find_first u terms2
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    64
    and T = Term.fastype_of u
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    65
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 43597
diff changeset
    66
    val t'' = Data.mk_bal (Data.mk_sum T (u::terms1'), Data.mk_sum T (u::terms2'))
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    67
    val reshape =
70315
2f9623aa1eeb proper context;
wenzelm
parents: 51717
diff changeset
    68
      Goal.prove ctxt' [] [] (Data.mk_eq (t', t'')) (K (Data.norm_tac ctxt))
45270
d5b5c9259afd fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents: 43597
diff changeset
    69
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    70
  in
70315
2f9623aa1eeb proper context;
wenzelm
parents: 51717
diff changeset
    71
    SOME (singleton (Variable.export ctxt' ctxt) (Data.simplify_meta_eq ctxt' simp_th reshape))
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    72
  end
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    73
  (* FIXME avoid handling of generic exceptions *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    74
  handle TERM _ => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    75
       | TYPE _ => NONE;   (*Typically (if thy doesn't include Numeral)
20114
a1bb4bc68ff3 prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents: 20044
diff changeset
    76
                             Undeclared type constructor "Numeral.bin"*)
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    77
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    78
end;