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