src/HOL/eqrule_HOL_data.ML
author skalberg
Sun, 13 Feb 2005 17:15:14 +0100
changeset 15531 08c8dad8e399
parent 15481 fc075ae929e4
child 15570 8d8c70b41bab
permissions -rw-r--r--
Deleted Library.option type.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     1
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     2
(*  Title:      sys/eqrule_HOL_data.ML
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     3
    Author:     Lucas Dixon, University of Edinburgh
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     4
                lucas.dixon@ed.ac.uk
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     5
    Modified:   22 July 2004
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     6
    Created:    18 Feb 2004
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     7
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     8
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
     9
(*  DESCRIPTION:
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    10
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    11
    Data for equality rules in the logic
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    12
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    13
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    14
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    15
(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    16
structure HOL_EqRuleData : EQRULE_DATA =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    17
struct
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    18
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    19
val eq_reflection = thm "eq_reflection";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    20
val mp = thm "mp";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    21
val spec = thm "spec";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    22
val if_bool_eq_conj = thm "if_bool_eq_conj";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    23
val iffD1 = thm "iffD1";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    24
val conjunct2 = thm "conjunct2";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    25
val conjunct1 = thm "conjunct1";
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    26
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    27
fun mk_eq th = case concl_of th of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    28
        Const("==",_)$_$_       => SOME (th)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    29
    |   _$(Const("op =",_)$_$_) => SOME (th RS eq_reflection)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    30
    |   _ => NONE;
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    31
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    32
val tranformation_pairs =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    33
  [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    34
   ("All", [spec]), ("True", []), ("False", []),
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    35
   ("If", [if_bool_eq_conj RS iffD1])];
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    36
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    37
(*
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    38
val mk_atomize:      (string * thm list) list -> thm -> thm list
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    39
looks too specific to move it somewhere else
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    40
*)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    41
fun mk_atomize pairs =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    42
  let fun atoms th =
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    43
        (case Thm.concl_of th of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    44
           Const("Trueprop",_) $ p =>
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    45
             (case Term.head_of p of
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    46
                Const(a,_) =>
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    47
                  (case Library.assoc(pairs,a) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    48
                     SOME(rls) => flat (map atoms ([th] RL rls))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15481
diff changeset
    49
                   | NONE => [th])
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    50
              | _ => [th])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    51
         | _ => [th])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    52
  in atoms end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    53
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    54
val prep_meta_eq = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    55
    (mapfilter  
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    56
       mk_eq
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    57
       o (mk_atomize tranformation_pairs)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    58
       o Drule.gen_all 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    59
       o zero_var_indexes)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    60
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    61
end;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    62
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    63
structure EqRuleData = HOL_EqRuleData;
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    64
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    65
structure EQSubstTac = 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    66
  EQSubstTacFUN(structure EqRuleData = EqRuleData);
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    67
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents:
diff changeset
    68