src/Provers/Arith/extract_common_term.ML
author wenzelm
Tue, 13 Sep 2005 22:19:23 +0200
changeset 17339 ab97ccef124a
parent 17223 430edc6b7826
child 17412 e26cb20ef0cc
permissions -rw-r--r--
tuned Isar interfaces; tuned IsarThy.theorem_i;
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
    ID:         $Id$
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     4
    Copyright   2000  University of Cambridge
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     5
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     6
Extract common terms in balanced expressions:
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     7
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     8
     i + u + j ~~ i' + u + j'  ==  u + (i + j) ~~ u + (i' + j')
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
     9
     i + u     ~~ u            ==  u + i       ~~ u + 0
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    10
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    11
where ~~ is an appropriate balancing operation (e.g. =, <=, <, -) and 0 is a 
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    12
suitable identity for +.
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    13
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    14
This massaged formula is then simplified in a user-specified way.
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
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    17
signature EXTRACT_COMMON_TERM_DATA =
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    18
sig
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    19
  (*abstract syntax*)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 13484
diff changeset
    20
  val mk_sum: typ -> term list -> term
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    21
  val dest_sum: term -> term list
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    22
  val mk_bal: term * term -> term
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    23
  val dest_bal: term -> term * term
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    24
  val find_first: term -> term list -> term list
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    25
  (*proof tools*)
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    26
  val prove_conv: tactic list -> theory -> 
13484
d8f5d3391766 adhoc_freeze_vars;
wenzelm
parents: 10694
diff changeset
    27
                  thm list -> string list -> term * term -> thm option
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    28
  val norm_tac: simpset -> tactic                (*proves the result*)
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    29
  val simplify_meta_eq: simpset -> thm -> thm    (*simplifies the result*)
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    30
end;
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
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    33
functor ExtractCommonTermFun(Data: EXTRACT_COMMON_TERM_DATA):
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    34
  sig
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    35
  val proc: theory -> simpset -> term -> thm option
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    36
  end 
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    37
=
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    38
struct
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    39
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    40
(*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
    41
fun find_common (terms1,terms2) =
17223
430edc6b7826 curried_lookup/update;
wenzelm
parents: 16973
diff changeset
    42
  let val tab2 = fold (Termtab.curried_update o rpair ()) terms2 Termtab.empty
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    43
      fun seek [] = raise TERM("find_common", []) 
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    44
	| seek (u::terms) =
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    45
	      if Termtab.defined tab2 u then u
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    46
	      else seek terms
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    47
  in seek terms1 end;
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    48
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    49
(*the simplification procedure*)
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    50
fun proc thy ss t =
15027
d23887300b96 adapted type of simprocs;
wenzelm
parents: 14387
diff changeset
    51
  let
d23887300b96 adapted type of simprocs;
wenzelm
parents: 14387
diff changeset
    52
      val hyps = prems_of_ss ss;
d23887300b96 adapted type of simprocs;
wenzelm
parents: 14387
diff changeset
    53
      (*first freeze any Vars in the term to prevent flex-flex problems*)
13484
d8f5d3391766 adhoc_freeze_vars;
wenzelm
parents: 10694
diff changeset
    54
      val (t', xs) = Term.adhoc_freeze_vars t;
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    55
      val (t1,t2) = Data.dest_bal t' 
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    56
      val terms1 = Data.dest_sum t1
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    57
      and terms2 = Data.dest_sum t2
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    58
      val u = find_common (terms1,terms2)
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    59
      val terms1' = Data.find_first u terms1
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    60
      and terms2' = Data.find_first u terms2
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 13484
diff changeset
    61
      and T = Term.fastype_of u
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    62
      val reshape = 
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    63
	    Data.prove_conv [Data.norm_tac ss] thy hyps xs
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    64
	        (t', 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 13484
diff changeset
    65
		 Data.mk_bal (Data.mk_sum T (u::terms1'), 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 13484
diff changeset
    66
		              Data.mk_sum T (u::terms2')))
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    67
  in
16973
b2a894562b8f simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15570
diff changeset
    68
      Option.map (Data.simplify_meta_eq ss) reshape
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    69
  end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    70
  handle TERM _ => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    71
       | TYPE _ => NONE;   (*Typically (if thy doesn't include Numeral)
10694
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    72
			     Undeclared type constructor "Numeral.bin"*)
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    73
9a5d5df29e5c new simproc for cancelling common factors, etc.
paulson
parents:
diff changeset
    74
end;