TFL/thms.sml
author paulson
Wed, 21 May 1997 10:59:14 +0200
changeset 3273 114704740c86
parent 3245 241838c01caf
child 3302 404fe31fd8d2
permissions -rw-r--r--
Code tidying: removal of C combinator
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     1
structure Thms : Thms_sig =
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     2
struct
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
     3
   val WFREC_COROLLARY = get_thm WF_Rel.thy "tfl_wfrec"
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
     4
   val WF_INDUCTION_THM = get_thm WF_Rel.thy "tfl_wf_induct"
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2112
diff changeset
     5
   val CUT_LEMMA = get_thm WF_Rel.thy "tfl_cut_apply"
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     6
   val CUT_DEF = cut_def
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     7
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     8
   local val _ = goal HOL.thy "!P x. P x --> P (Eps P)"
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     9
         val _ = by (strip_tac 1)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    10
         val _ = by (rtac selectI 1)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    11
         val _ = by (assume_tac 1)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    12
   in val SELECT_AX = result()
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    13
   end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    15
  (*-------------------------------------------------------------------------
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    16
   *  A useful congruence rule
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    17
   *-------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    18
   local val [p1,p2] = goal HOL.thy "(M = N) ==> \
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    19
                          \ (!!x. ((x = N) ==> (f x = g x))) ==> \
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    20
                          \ (Let M f = Let N g)";
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    21
         val _ = by (simp_tac (HOL_ss addsimps[Let_def,p1]) 1);
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    22
         val _ = by (rtac p2 1);
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    23
         val _ = by (rtac refl 1);
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    24
   in val LET_CONG = result() RS eq_reflection
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    25
   end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    26
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    27
   val COND_CONG = if_cong RS eq_reflection;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    28
3273
114704740c86 Code tidying: removal of C combinator
paulson
parents: 3245
diff changeset
    29
   fun prove s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    30
3273
114704740c86 Code tidying: removal of C combinator
paulson
parents: 3245
diff changeset
    31
   val eqT       = prove"(x = True) --> x"
114704740c86 Code tidying: removal of C combinator
paulson
parents: 3245
diff changeset
    32
   val rev_eq_mp = prove"(x = y) --> y --> x"
114704740c86 Code tidying: removal of C combinator
paulson
parents: 3245
diff changeset
    33
   val simp_thm  = prove"(x-->y) --> (x = x') --> (x' --> y)"
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    34
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    35
end;