src/Sequents/LK0.thy
author wenzelm
Sat, 22 Mar 2014 18:19:57 +0100
changeset 56254 a2dd9200854d
parent 55380 4de48353034e
child 58860 fee7cfa69c50
permissions -rw-r--r--
more antiquotations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 39159
diff changeset
     1
(*  Title:      Sequents/LK0.thy
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
     4
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
     5
There may be printing problems if a seqent is in expanded normal form
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
     6
(eta-expanded, beta-contracted).
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
     7
*)
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
     8
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
     9
header {* Classical First-Order Sequent Calculus *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    10
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    11
theory LK0
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    12
imports Sequents
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    13
begin
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    14
55380
4de48353034e prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents: 55233
diff changeset
    15
class "term"
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35417
diff changeset
    16
default_sort "term"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    17
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    18
consts
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    19
21524
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 21428
diff changeset
    20
  Trueprop       :: "two_seqi"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    21
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    22
  True         :: o
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    23
  False        :: o
22894
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    24
  equal        :: "['a,'a] => o"     (infixl "=" 50)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    25
  Not          :: "o => o"           ("~ _" [40] 40)
22894
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    26
  conj         :: "[o,o] => o"       (infixr "&" 35)
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    27
  disj         :: "[o,o] => o"       (infixr "|" 30)
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    28
  imp          :: "[o,o] => o"       (infixr "-->" 25)
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    29
  iff          :: "[o,o] => o"       (infixr "<->" 25)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    30
  The          :: "('a => o) => 'a"  (binder "THE " 10)
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    31
  All          :: "('a => o) => o"   (binder "ALL " 10)
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    32
  Ex           :: "('a => o) => o"   (binder "EX " 10)
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    33
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    34
syntax
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    35
 "_Trueprop"    :: "two_seqe" ("((_)/ |- (_))" [6,6] 5)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    36
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51309
diff changeset
    37
parse_translation {* [(@{syntax_const "_Trueprop"}, K (two_seq_tr @{const_syntax Trueprop}))] *}
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51309
diff changeset
    38
print_translation {* [(@{const_syntax Trueprop}, K (two_seq_tr' @{syntax_const "_Trueprop"}))] *}
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    39
22894
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    40
abbreviation
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    41
  not_equal  (infixl "~=" 50) where
619b270607ac eliminated unnamed infixes;
wenzelm
parents: 21588
diff changeset
    42
  "x ~= y == ~ (x = y)"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    43
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    44
notation (xsymbols)
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    45
  Not  ("\<not> _" [40] 40) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    46
  conj  (infixr "\<and>" 35) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    47
  disj  (infixr "\<or>" 30) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    48
  imp  (infixr "\<longrightarrow>" 25) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    49
  iff  (infixr "\<longleftrightarrow>" 25) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    50
  All  (binder "\<forall>" 10) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    51
  Ex  (binder "\<exists>" 10) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    52
  not_equal  (infixl "\<noteq>" 50)
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    53
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    54
notation (HTML output)
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    55
  Not  ("\<not> _" [40] 40) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    56
  conj  (infixr "\<and>" 35) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    57
  disj  (infixr "\<or>" 30) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    58
  All  (binder "\<forall>" 10) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    59
  Ex  (binder "\<exists>" 10) and
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35113
diff changeset
    60
  not_equal  (infixl "\<noteq>" 50)
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    61
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    62
axiomatization where
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    63
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    64
  (*Structural rules: contraction, thinning, exchange [Soren Heilmann] *)
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    65
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    66
  contRS: "$H |- $E, $S, $S, $F ==> $H |- $E, $S, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    67
  contLS: "$H, $S, $S, $G |- $E ==> $H, $S, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    68
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    69
  thinRS: "$H |- $E, $F ==> $H |- $E, $S, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    70
  thinLS: "$H, $G |- $E ==> $H, $S, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    71
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    72
  exchRS: "$H |- $E, $R, $S, $F ==> $H |- $E, $S, $R, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    73
  exchLS: "$H, $R, $S, $G |- $E ==> $H, $S, $R, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    74
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    75
  cut:   "[| $H |- $E, P;  $H, P |- $E |] ==> $H |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    76
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    77
  (*Propositional rules*)
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    78
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    79
  basic: "$H, P, $G |- $E, P, $F" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    80
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    81
  conjR: "[| $H|- $E, P, $F;  $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    82
  conjL: "$H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    83
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    84
  disjR: "$H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    85
  disjL: "[| $H, P, $G |- $E;  $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    86
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    87
  impR:  "$H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    88
  impL:  "[| $H,$G |- $E,P;  $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    89
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    90
  notR:  "$H, P |- $E, $F ==> $H |- $E, ~P, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    91
  notL:  "$H, $G |- $E, P ==> $H, ~P, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    92
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    93
  FalseL: "$H, False, $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    94
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    95
  True_def: "True == False-->False" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
    96
  iff_def:  "P<->Q == (P-->Q) & (Q-->P)"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    97
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
    98
axiomatization where
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
    99
  (*Quantifiers*)
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   100
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   101
  allR:  "(!!x.$H |- $E, P(x), $F) ==> $H |- $E, ALL x. P(x), $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   102
  allL:  "$H, P(x), $G, ALL x. P(x) |- $E ==> $H, ALL x. P(x), $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   103
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   104
  exR:   "$H |- $E, P(x), $F, EX x. P(x) ==> $H |- $E, EX x. P(x), $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   105
  exL:   "(!!x.$H, P(x), $G |- $E) ==> $H, EX x. P(x), $G |- $E" and
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   106
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   107
  (*Equality*)
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   108
  refl:  "$H |- $E, a=a, $F" and
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   109
  subst: "\<And>G H E. $H(a), $G(a) |- $E(a) ==> $H(b), a=b, $G(b) |- $E(b)"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   110
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   111
  (* Reflection *)
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   112
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   113
axiomatization where
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   114
  eq_reflection:  "|- x=y ==> (x==y)" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
   115
  iff_reflection: "|- P<->Q ==> (P==Q)"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   116
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   117
  (*Descriptions*)
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   118
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   119
axiomatization where
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 16019
diff changeset
   120
  The: "[| $H |- $E, P(a), $F;  !!x.$H, P(x) |- $E, x=a, $F |] ==>
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   121
          $H |- $E, P(THE x. P(x)), $F"
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   122
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   123
definition If :: "[o, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10)
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 45602
diff changeset
   124
  where "If(P,x,y) == THE z::'a. (P --> z=x) & (~P --> z=y)"
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   125
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   126
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   127
(** Structural Rules on formulas **)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   128
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   129
(*contraction*)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   130
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   131
lemma contR: "$H |- $E, P, P, $F ==> $H |- $E, P, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   132
  by (rule contRS)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   133
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   134
lemma contL: "$H, P, P, $G |- $E ==> $H, P, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   135
  by (rule contLS)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   136
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   137
(*thinning*)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   138
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   139
lemma thinR: "$H |- $E, $F ==> $H |- $E, P, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   140
  by (rule thinRS)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   141
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   142
lemma thinL: "$H, $G |- $E ==> $H, P, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   143
  by (rule thinLS)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   144
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   145
(*exchange*)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   146
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   147
lemma exchR: "$H |- $E, Q, P, $F ==> $H |- $E, P, Q, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   148
  by (rule exchRS)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   149
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   150
lemma exchL: "$H, Q, P, $G |- $E ==> $H, P, Q, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   151
  by (rule exchLS)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   152
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   153
ML {*
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   154
(*Cut and thin, replacing the right-side formula*)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27146
diff changeset
   155
fun cutR_tac ctxt s i =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27208
diff changeset
   156
  res_inst_tac ctxt [(("P", 0), s) ] @{thm cut} i  THEN  rtac @{thm thinR} i
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   157
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   158
(*Cut and thin, replacing the left-side formula*)
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 27146
diff changeset
   159
fun cutL_tac ctxt s i =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27208
diff changeset
   160
  res_inst_tac ctxt [(("P", 0), s)] @{thm cut} i  THEN  rtac @{thm thinL} (i+1)
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   161
*}
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   162
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   163
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   164
(** If-and-only-if rules **)
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   165
lemma iffR:
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   166
    "[| $H,P |- $E,Q,$F;  $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   167
  apply (unfold iff_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   168
  apply (assumption | rule conjR impR)+
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   169
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   170
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   171
lemma iffL:
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   172
    "[| $H,$G |- $E,P,Q;  $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   173
  apply (unfold iff_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   174
  apply (assumption | rule conjL impL basic)+
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   175
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   176
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   177
lemma iff_refl: "$H |- $E, (P <-> P), $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   178
  apply (rule iffR basic)+
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   179
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   180
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   181
lemma TrueR: "$H |- $E, True, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   182
  apply (unfold True_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   183
  apply (rule impR)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   184
  apply (rule basic)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   185
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   186
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   187
(*Descriptions*)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   188
lemma the_equality:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   189
  assumes p1: "$H |- $E, P(a), $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   190
    and p2: "!!x. $H, P(x) |- $E, x=a, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   191
  shows "$H |- $E, (THE x. P(x)) = a, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   192
  apply (rule cut)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   193
   apply (rule_tac [2] p2)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   194
  apply (rule The, rule thinR, rule exchRS, rule p1)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   195
  apply (rule thinR, rule exchRS, rule p2)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   196
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   197
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   198
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   199
(** Weakened quantifier rules.  Incomplete, they let the search terminate.**)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   200
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   201
lemma allL_thin: "$H, P(x), $G |- $E ==> $H, ALL x. P(x), $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   202
  apply (rule allL)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   203
  apply (erule thinL)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   204
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   205
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   206
lemma exR_thin: "$H |- $E, P(x), $F ==> $H |- $E, EX x. P(x), $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   207
  apply (rule exR)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   208
  apply (erule thinR)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   209
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   210
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   211
(*The rules of LK*)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   212
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   213
lemmas [safe] =
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   214
  iffR iffL
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   215
  notR notL
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   216
  impR impL
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   217
  disjR disjL
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   218
  conjR conjL
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   219
  FalseL TrueR
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   220
  refl basic
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   221
ML {* val prop_pack = Cla.get_pack @{context} *}
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   222
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   223
lemmas [safe] = exL allR
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   224
lemmas [unsafe] = the_equality exR_thin allL_thin
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   225
ML {* val LK_pack = Cla.get_pack @{context} *}
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   226
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   227
ML {*
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   228
  val LK_dup_pack =
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   229
    Cla.put_pack prop_pack @{context}
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   230
    |> fold_rev Cla.add_safe @{thms allR exL}
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   231
    |> fold_rev Cla.add_unsafe @{thms allL exR the_equality}
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   232
    |> Cla.get_pack;
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   233
*}
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   234
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   235
method_setup fast_prop =
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   236
  {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack prop_pack ctxt))) *}
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   237
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   238
method_setup fast_dup =
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   239
  {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack LK_dup_pack ctxt))) *}
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   240
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   241
method_setup best_dup =
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   242
  {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack LK_dup_pack ctxt))) *}
7093
b2ee0e5d1a7f renamed theory LK to LK0
paulson
parents:
diff changeset
   243
55233
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   244
method_setup lem = {*
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   245
  Attrib.thm >> (fn th => fn _ =>
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   246
    SIMPLE_METHOD' (fn i =>
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   247
      rtac (@{thm thinR} RS @{thm cut}) i THEN
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   248
      REPEAT (rtac @{thm thinL} i) THEN
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   249
      rtac th i))
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   250
*}
3229614ca9c5 method_setup "lem";
wenzelm
parents: 55228
diff changeset
   251
7118
ee384c7b7416 adding missing declarations for the <<...>> notation
paulson
parents: 7093
diff changeset
   252
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   253
lemma mp_R:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   254
  assumes major: "$H |- $E, $F, P --> Q"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   255
    and minor: "$H |- $E, $F, P"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   256
  shows "$H |- $E, Q, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   257
  apply (rule thinRS [THEN cut], rule major)
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   258
  apply step
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   259
  apply (rule thinR, rule minor)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   260
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   261
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   262
lemma mp_L:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   263
  assumes major: "$H, $G |- $E, P --> Q"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   264
    and minor: "$H, $G, Q |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   265
  shows "$H, P, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   266
  apply (rule thinL [THEN cut], rule major)
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   267
  apply step
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   268
  apply (rule thinL, rule minor)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   269
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   270
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   271
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   272
(** Two rules to generate left- and right- rules from implications **)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   273
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   274
lemma R_of_imp:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   275
  assumes major: "|- P --> Q"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   276
    and minor: "$H |- $E, $F, P"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   277
  shows "$H |- $E, Q, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   278
  apply (rule mp_R)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   279
   apply (rule_tac [2] minor)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   280
  apply (rule thinRS, rule major [THEN thinLS])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   281
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   282
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   283
lemma L_of_imp:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   284
  assumes major: "|- P --> Q"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   285
    and minor: "$H, $G, Q |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   286
  shows "$H, P, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   287
  apply (rule mp_L)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   288
   apply (rule_tac [2] minor)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   289
  apply (rule thinRS, rule major [THEN thinLS])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   290
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   291
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   292
(*Can be used to create implications in a subgoal*)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   293
lemma backwards_impR:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   294
  assumes prem: "$H, $G |- $E, $F, P --> Q"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   295
  shows "$H, P, $G |- $E, Q, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   296
  apply (rule mp_L)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   297
   apply (rule_tac [2] basic)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   298
  apply (rule thinR, rule prem)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   299
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   300
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   301
lemma conjunct1: "|-P&Q ==> |-P"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   302
  apply (erule thinR [THEN cut])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   303
  apply fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   304
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   305
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   306
lemma conjunct2: "|-P&Q ==> |-Q"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   307
  apply (erule thinR [THEN cut])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   308
  apply fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   309
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   310
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   311
lemma spec: "|- (ALL x. P(x)) ==> |- P(x)"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   312
  apply (erule thinR [THEN cut])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   313
  apply fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   314
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   315
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   316
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   317
(** Equality **)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   318
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   319
lemma sym: "|- a=b --> b=a"
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   320
  by (safe add!: subst)
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   321
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   322
lemma trans: "|- a=b --> b=c --> a=c"
55228
901a6696cdd8 misc tuning and modernization;
wenzelm
parents: 52143
diff changeset
   323
  by (safe add!: subst)
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   324
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   325
(* Symmetry of equality in hypotheses *)
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42814
diff changeset
   326
lemmas symL = sym [THEN L_of_imp]
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   327
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   328
(* Symmetry of equality in hypotheses *)
45602
2a858377c3d2 eliminated obsolete "standard";
wenzelm
parents: 42814
diff changeset
   329
lemmas symR = sym [THEN R_of_imp]
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   330
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   331
lemma transR: "[| $H|- $E, $F, a=b;  $H|- $E, $F, b=c |] ==> $H|- $E, a=c, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   332
  by (rule trans [THEN R_of_imp, THEN mp_R])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   333
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   334
(* Two theorms for rewriting only one instance of a definition:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   335
   the first for definitions of formulae and the second for terms *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   336
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   337
lemma def_imp_iff: "(A == B) ==> |- A <-> B"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   338
  apply unfold
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   339
  apply (rule iff_refl)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   340
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   341
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   342
lemma meta_eq_to_obj_eq: "(A == B) ==> |- A = B"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   343
  apply unfold
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   344
  apply (rule refl)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   345
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   346
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   347
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   348
(** if-then-else rules **)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   349
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   350
lemma if_True: "|- (if True then x else y) = x"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   351
  unfolding If_def by fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   352
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   353
lemma if_False: "|- (if False then x else y) = y"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   354
  unfolding If_def by fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   355
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   356
lemma if_P: "|- P ==> |- (if P then x else y) = x"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   357
  apply (unfold If_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   358
  apply (erule thinR [THEN cut])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   359
  apply fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   360
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   361
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   362
lemma if_not_P: "|- ~P ==> |- (if P then x else y) = y";
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   363
  apply (unfold If_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   364
  apply (erule thinR [THEN cut])
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   365
  apply fast
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   366
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   367
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   368
end