src/LCF/LCF_lemmas.ML
author wenzelm
Mon, 17 Oct 2005 23:10:13 +0200
changeset 17876 b9c92f384109
parent 17248 81bf91654e73
child 19755 90f80de04c46
permissions -rw-r--r--
change_claset/simpset;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17248
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     1
(*  Title:      LCF/lcf.ML
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     3
    Author:     Tobias Nipkow
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     5
*)
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     6
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     7
(* Standard abbreviations *)
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     8
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
     9
val rstac = resolve_tac;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    10
fun stac th = rtac(th RS sym RS subst);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    11
fun sstac ths = EVERY' (map stac ths);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    12
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    13
fun strip_tac i = REPEAT(rstac [impI,allI] i);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    14
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    15
val eq_imp_less1 = prove_goal (the_context ()) "x=y ==> x << y"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    16
        (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct1) 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    17
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    18
val eq_imp_less2 = prove_goal (the_context ()) "x=y ==> y << x"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    19
        (fn prems => [rtac (rewrite_rule[eq_def](hd prems) RS conjunct2) 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    20
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    21
val less_refl = refl RS eq_imp_less1;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    22
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    23
val less_anti_sym = prove_goal (the_context ()) "[| x << y; y << x |] ==> x=y"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    24
        (fn prems => [rewtac eq_def,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    25
                      REPEAT(rstac(conjI::prems)1)]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    26
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    27
val ext = prove_goal (the_context ())
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    28
        "(!!x::'a::cpo. f(x)=(g(x)::'b::cpo)) ==> (%x. f(x))=(%x. g(x))"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    29
        (fn [prem] => [REPEAT(rstac[less_anti_sym, less_ext, allI,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    30
                                    prem RS eq_imp_less1,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    31
                                    prem RS eq_imp_less2]1)]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    32
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    33
val cong = prove_goal (the_context ()) "[| f=g; x=y |] ==> f(x)=g(y)"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    34
        (fn prems => [cut_facts_tac prems 1, etac subst 1, etac subst 1,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    35
                      rtac refl 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    36
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    37
val less_ap_term = less_refl RS mono;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    38
val less_ap_thm = less_refl RSN (2,mono);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    39
val ap_term = refl RS cong;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    40
val ap_thm = refl RSN (2,cong);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    41
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    42
val UU_abs = prove_goal (the_context ()) "(%x::'a::cpo. UU) = UU"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    43
        (fn _ => [rtac less_anti_sym 1, rtac minimal 2,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    44
                  rtac less_ext 1, rtac allI 1, rtac minimal 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    45
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    46
val UU_app = UU_abs RS sym RS ap_thm;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    47
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    48
val less_UU = prove_goal (the_context ()) "x << UU ==> x=UU"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    49
        (fn prems=> [rtac less_anti_sym 1,rstac prems 1,rtac minimal 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    50
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    51
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    52
val tr_induct = prove_goal (the_context ()) "[| P(UU); P(TT); P(FF) |] ==> ALL b. P(b)"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    53
        (fn prems => [rtac allI 1, rtac mp 1,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    54
                      res_inst_tac[("p","b")]tr_cases 2,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    55
                      fast_tac (FOL_cs addIs prems) 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    56
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    57
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    58
val Contrapos = prove_goal (the_context ()) "(A ==> B) ==> (~B ==> ~A)"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    59
        (fn prems => [rtac notI 1, rtac notE 1, rstac prems 1,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    60
                      rstac prems 1, atac 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    61
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    62
val not_less_imp_not_eq1 = eq_imp_less1 COMP Contrapos;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    63
val not_less_imp_not_eq2 = eq_imp_less2 COMP Contrapos;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    64
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    65
val not_UU_eq_TT = not_TT_less_UU RS not_less_imp_not_eq2;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    66
val not_UU_eq_FF = not_FF_less_UU RS not_less_imp_not_eq2;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    67
val not_TT_eq_UU = not_TT_less_UU RS not_less_imp_not_eq1;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    68
val not_TT_eq_FF = not_TT_less_FF RS not_less_imp_not_eq1;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    69
val not_FF_eq_UU = not_FF_less_UU RS not_less_imp_not_eq1;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    70
val not_FF_eq_TT = not_FF_less_TT RS not_less_imp_not_eq1;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    71
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    72
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    73
val COND_cases_iff = (prove_goal (the_context ())
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    74
  "ALL b. P(b=>x|y) <-> (b=UU-->P(UU)) & (b=TT-->P(x)) & (b=FF-->P(y))"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    75
        (fn _ => [cut_facts_tac [not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    76
                                 not_TT_eq_FF,not_FF_eq_UU,not_FF_eq_TT]1,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    77
                  rtac tr_induct 1, stac COND_UU 1, stac COND_TT 2,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    78
                  stac COND_FF 3, REPEAT(fast_tac FOL_cs 1)]))  RS spec;
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    79
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    80
val lemma = prove_goal (the_context ()) "A<->B ==> B ==> A"
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    81
        (fn prems => [cut_facts_tac prems 1, rewtac iff_def,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    82
                      fast_tac FOL_cs 1]);
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    83
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    84
val COND_cases = conjI RSN (2,conjI RS (COND_cases_iff RS lemma));
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    85
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    86
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    87
val LCF_ss = FOL_ss addsimps
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    88
        [minimal,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    89
         UU_app, UU_app RS ap_thm, UU_app RS ap_thm RS ap_thm,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    90
         not_TT_less_FF,not_FF_less_TT,not_TT_less_UU,not_FF_less_UU,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    91
         not_UU_eq_TT,not_UU_eq_FF,not_TT_eq_UU,not_TT_eq_FF,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    92
         not_FF_eq_UU,not_FF_eq_TT,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    93
         COND_UU,COND_TT,COND_FF,
81bf91654e73 converted to Isar theory format;
wenzelm
parents:
diff changeset
    94
         surj_pairing,FST,SND];
17876
b9c92f384109 change_claset/simpset;
wenzelm
parents: 17248
diff changeset
    95
b9c92f384109 change_claset/simpset;
wenzelm
parents: 17248
diff changeset
    96
change_simpset (fn _ => LCF_ss);