src/FOL/IFOL.thy
author wenzelm
Sat, 20 May 2006 23:36:49 +0200
changeset 19683 3620e494cef2
parent 19656 09be06943252
child 19756 61c4117345c6
permissions -rw-r--r--
removed obsolete 'finalconsts';
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1268
f6ef556b3ede corrected title
clasohm
parents: 928
diff changeset
     1
(*  Title:      FOL/IFOL.thy
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     2
    ID:         $Id$
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
     3
    Author:     Lawrence C Paulson and Markus Wenzel
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
     4
*)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     5
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
     6
header {* Intuitionistic first-order logic *}
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     7
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15377
diff changeset
     8
theory IFOL
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15377
diff changeset
     9
imports Pure
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 16121
diff changeset
    10
uses ("IFOL_lemmas.ML") ("fologic.ML") ("hypsubstdata.ML") ("intprover.ML")
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15377
diff changeset
    11
begin
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    12
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    14
subsection {* Syntax and axiomatic basis *}
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    15
3906
5ae0e1324c56 global;
wenzelm
parents: 3835
diff changeset
    16
global
5ae0e1324c56 global;
wenzelm
parents: 3835
diff changeset
    17
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14565
diff changeset
    18
classes "term"
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    19
defaultsort "term"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    21
typedecl o
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    22
11747
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
    23
judgment
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
    24
  Trueprop      :: "o => prop"                  ("(_)" 5)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
11747
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
    26
consts
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    27
  True          :: o
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    28
  False         :: o
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    29
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    30
  (* Connectives *)
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    31
17276
3bb24e8b2cb2 avoid old-style infixes;
wenzelm
parents: 16417
diff changeset
    32
  "op ="        :: "['a, 'a] => o"              (infixl "=" 50)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    33
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    34
  Not           :: "o => o"                     ("~ _" [40] 40)
17276
3bb24e8b2cb2 avoid old-style infixes;
wenzelm
parents: 16417
diff changeset
    35
  "op &"        :: "[o, o] => o"                (infixr "&" 35)
3bb24e8b2cb2 avoid old-style infixes;
wenzelm
parents: 16417
diff changeset
    36
  "op |"        :: "[o, o] => o"                (infixr "|" 30)
3bb24e8b2cb2 avoid old-style infixes;
wenzelm
parents: 16417
diff changeset
    37
  "op -->"      :: "[o, o] => o"                (infixr "-->" 25)
3bb24e8b2cb2 avoid old-style infixes;
wenzelm
parents: 16417
diff changeset
    38
  "op <->"      :: "[o, o] => o"                (infixr "<->" 25)
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    39
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    40
  (* Quantifiers *)
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    41
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    42
  All           :: "('a => o) => o"             (binder "ALL " 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    43
  Ex            :: "('a => o) => o"             (binder "EX " 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    44
  Ex1           :: "('a => o) => o"             (binder "EX! " 10)
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    45
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19120
diff changeset
    47
abbreviation
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
    48
  not_equal     :: "['a, 'a] => o"              (infixl "~=" 50)
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
    49
  "x ~= y == ~ (x = y)"
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    50
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19380
diff changeset
    51
const_syntax (xsymbols)
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19380
diff changeset
    52
  not_equal  (infixl "\<noteq>" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19120
diff changeset
    53
19656
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19380
diff changeset
    54
const_syntax (HTML output)
09be06943252 tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents: 19380
diff changeset
    55
  not_equal  (infixl "\<noteq>" 50)
19363
667b5ea637dd refined 'abbreviation';
wenzelm
parents: 19120
diff changeset
    56
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 12019
diff changeset
    57
syntax (xsymbols)
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    58
  Not           :: "o => o"                     ("\<not> _" [40] 40)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    59
  "op &"        :: "[o, o] => o"                (infixr "\<and>" 35)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    60
  "op |"        :: "[o, o] => o"                (infixr "\<or>" 30)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    61
  "ALL "        :: "[idts, o] => o"             ("(3\<forall>_./ _)" [0, 10] 10)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    62
  "EX "         :: "[idts, o] => o"             ("(3\<exists>_./ _)" [0, 10] 10)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    63
  "EX! "        :: "[idts, o] => o"             ("(3\<exists>!_./ _)" [0, 10] 10)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    64
  "op -->"      :: "[o, o] => o"                (infixr "\<longrightarrow>" 25)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    65
  "op <->"      :: "[o, o] => o"                (infixr "\<longleftrightarrow>" 25)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    66
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6027
diff changeset
    67
syntax (HTML output)
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    68
  Not           :: "o => o"                     ("\<not> _" [40] 40)
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14236
diff changeset
    69
  "op &"        :: "[o, o] => o"                (infixr "\<and>" 35)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14236
diff changeset
    70
  "op |"        :: "[o, o] => o"                (infixr "\<or>" 30)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14236
diff changeset
    71
  "ALL "        :: "[idts, o] => o"             ("(3\<forall>_./ _)" [0, 10] 10)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14236
diff changeset
    72
  "EX "         :: "[idts, o] => o"             ("(3\<exists>_./ _)" [0, 10] 10)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14236
diff changeset
    73
  "EX! "        :: "[idts, o] => o"             ("(3\<exists>!_./ _)" [0, 10] 10)
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6027
diff changeset
    74
7d5cbd5819a0 HTML output;
wenzelm
parents: 6027
diff changeset
    75
3932
wenzelm
parents: 3906
diff changeset
    76
local
3906
5ae0e1324c56 global;
wenzelm
parents: 3835
diff changeset
    77
14236
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    78
finalconsts
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    79
  False All Ex
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    80
  "op ="
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    81
  "op &"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    82
  "op |"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    83
  "op -->"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
    84
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    85
axioms
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    87
  (* Equality *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    89
  refl:         "a=a"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    91
  (* Propositional logic *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    93
  conjI:        "[| P;  Q |] ==> P&Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    94
  conjunct1:    "P&Q ==> P"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    95
  conjunct2:    "P&Q ==> Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    97
  disjI1:       "P ==> P|Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    98
  disjI2:       "Q ==> P|Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    99
  disjE:        "[| P|Q;  P ==> R;  Q ==> R |] ==> R"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   101
  impI:         "(P ==> Q) ==> P-->Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   102
  mp:           "[| P-->Q;  P |] ==> Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   104
  FalseE:       "False ==> P"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   105
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   106
  (* Quantifiers *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   108
  allI:         "(!!x. P(x)) ==> (ALL x. P(x))"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   109
  spec:         "(ALL x. P(x)) ==> P(x)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   111
  exI:          "P(x) ==> (EX x. P(x))"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   112
  exE:          "[| EX x. P(x);  !!x. P(x) ==> R |] ==> R"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
  (* Reflection *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   116
  eq_reflection:  "(x=y)   ==> (x==y)"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   117
  iff_reflection: "(P<->Q) ==> (P==Q)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   118
4092
9faf228771dc added simpset thy_data;
wenzelm
parents: 3932
diff changeset
   119
15377
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   120
text{*Thanks to Stephan Merz*}
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   121
theorem subst:
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   122
  assumes eq: "a = b" and p: "P(a)"
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   123
  shows "P(b)"
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   124
proof -
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   125
  from eq have meta: "a \<equiv> b"
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   126
    by (rule eq_reflection)
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   127
  from p show ?thesis
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   128
    by (unfold meta)
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   129
qed
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   130
3d99eea28a9b proof of subst by S Merz
paulson
parents: 14854
diff changeset
   131
14236
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   132
defs
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   133
  (* Definitions *)
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   134
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   135
  True_def:     "True  == False-->False"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   136
  not_def:      "~P    == P-->False"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   137
  iff_def:      "P<->Q == (P-->Q) & (Q-->P)"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   138
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   139
  (* Unique existence *)
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   140
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   141
  ex1_def:      "Ex1(P) == EX x. P(x) & (ALL y. P(y) --> y=x)"
c73d62ce9d1c partial conversion to Isar scripts
paulson
parents: 13779
diff changeset
   142
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   143
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   144
subsection {* Lemmas and proof tools *}
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   145
9886
897d6602cbfb updated setup;
wenzelm
parents: 9526
diff changeset
   146
use "IFOL_lemmas.ML"
11734
7a21bf539412 declare impE iffD1 iffD2 ad elim of Pure;
wenzelm
parents: 11677
diff changeset
   147
18481
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   148
ML {*
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   149
structure ProjectRule = ProjectRuleFun
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   150
(struct
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   151
  val conjunct1 = thm "conjunct1";
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   152
  val conjunct2 = thm "conjunct2";
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   153
  val mp = thm "mp";
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   154
end)
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   155
*}
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   156
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   157
use "fologic.ML"
9886
897d6602cbfb updated setup;
wenzelm
parents: 9526
diff changeset
   158
use "hypsubstdata.ML"
897d6602cbfb updated setup;
wenzelm
parents: 9526
diff changeset
   159
setup hypsubst_setup
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   160
use "intprover.ML"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   161
4092
9faf228771dc added simpset thy_data;
wenzelm
parents: 3932
diff changeset
   162
12875
wenzelm
parents: 12662
diff changeset
   163
subsection {* Intuitionistic Reasoning *}
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   164
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   165
lemma impE':
12937
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   166
  assumes 1: "P --> Q"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   167
    and 2: "Q ==> R"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   168
    and 3: "P --> Q ==> P"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   169
  shows R
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   170
proof -
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   171
  from 3 and 1 have P .
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   172
  with 1 have Q by (rule impE)
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   173
  with 2 show R .
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   174
qed
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   175
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   176
lemma allE':
12937
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   177
  assumes 1: "ALL x. P(x)"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   178
    and 2: "P(x) ==> ALL x. P(x) ==> Q"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   179
  shows Q
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   180
proof -
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   181
  from 1 have "P(x)" by (rule spec)
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   182
  from this and 1 show Q by (rule 2)
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   183
qed
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   184
12937
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   185
lemma notE':
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   186
  assumes 1: "~ P"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   187
    and 2: "~ P ==> P"
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   188
  shows R
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   189
proof -
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   190
  from 2 and 1 have P .
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   191
  with 1 show R by (rule notE)
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   192
qed
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   193
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   194
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   195
  and [Pure.intro!] = iffI conjI impI TrueI notI allI refl
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   196
  and [Pure.elim 2] = allE notE' impE'
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   197
  and [Pure.intro] = exI disjI2 disjI1
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   198
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18523
diff changeset
   199
setup {* ContextRules.addSWrapper (fn tac => hyp_subst_tac ORELSE' tac) *}
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   200
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   201
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   202
lemma iff_not_sym: "~ (Q <-> P) ==> ~ (P <-> Q)"
17591
33d409318266 renamed rules to iprover
nipkow
parents: 17276
diff changeset
   203
  by iprover
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   204
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   205
lemmas [sym] = sym iff_sym not_sym iff_not_sym
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   206
  and [Pure.elim?] = iffD1 iffD2 impE
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   207
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   208
13435
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   209
lemma eq_commute: "a=b <-> b=a"
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   210
apply (rule iffI) 
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   211
apply (erule sym)+
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   212
done
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   213
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   214
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   215
subsection {* Atomizing meta-level rules *}
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   216
11747
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
   217
lemma atomize_all [atomize]: "(!!x. P(x)) == Trueprop (ALL x. P(x))"
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   218
proof
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   219
  assume "!!x. P(x)"
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   220
  show "ALL x. P(x)" ..
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   221
next
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   222
  assume "ALL x. P(x)"
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   223
  thus "!!x. P(x)" ..
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   224
qed
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   225
11747
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
   226
lemma atomize_imp [atomize]: "(A ==> B) == Trueprop (A --> B)"
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   227
proof
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   228
  assume "A ==> B"
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   229
  thus "A --> B" ..
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   230
next
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   231
  assume "A --> B" and A
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   232
  thus B by (rule mp)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   233
qed
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   234
11747
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
   235
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)"
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   236
proof
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   237
  assume "x == y"
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   238
  show "x = y" by (unfold prems) (rule refl)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   239
next
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   240
  assume "x = y"
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   241
  thus "x == y" by (rule eq_reflection)
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   242
qed
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   243
18813
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   244
lemma atomize_iff [atomize]: "(A == B) == Trueprop (A <-> B)"
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   245
proof
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   246
  assume "A == B"
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   247
  show "A <-> B" by (unfold prems) (rule iff_refl)
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   248
next
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   249
  assume "A <-> B"
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   250
  thus "A == B" by (rule iff_reflection)
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   251
qed
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   252
12875
wenzelm
parents: 12662
diff changeset
   253
lemma atomize_conj [atomize]:
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   254
  includes meta_conjunction_syntax
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   255
  shows "(A && B) == Trueprop (A & B)"
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   256
proof
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   257
  assume conj: "A && B"
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   258
  show "A & B"
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   259
  proof (rule conjI)
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   260
    from conj show A by (rule conjunctionD1)
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   261
    from conj show B by (rule conjunctionD2)
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   262
  qed
11953
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   263
next
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   264
  assume conj: "A & B"
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   265
  show "A && B"
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   266
  proof -
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   267
    from conj show A ..
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   268
    from conj show B ..
11953
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   269
  qed
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   270
qed
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   271
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   272
lemmas [symmetric, rulify] = atomize_all atomize_imp
18861
38269ef5175a declare defn rules;
wenzelm
parents: 18813
diff changeset
   273
  and [symmetric, defn] = atomize_all atomize_imp atomize_eq atomize_iff
11771
b7b100a2de1d moved rulify to ObjectLogic;
wenzelm
parents: 11747
diff changeset
   274
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   275
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   276
subsection {* Calculational rules *}
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   277
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   278
lemma forw_subst: "a = b ==> P(b) ==> P(a)"
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   279
  by (rule ssubst)
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   280
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   281
lemma back_subst: "P(a) ==> a = b ==> P(b)"
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   282
  by (rule subst)
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   283
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   284
text {*
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   285
  Note that this list of rules is in reverse order of priorities.
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   286
*}
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   287
12019
abe9b7c6016e transitive declared in Pure;
wenzelm
parents: 11976
diff changeset
   288
lemmas basic_trans_rules [trans] =
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   289
  forw_subst
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   290
  back_subst
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   291
  rev_mp
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   292
  mp
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   293
  trans
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   294
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   295
subsection {* ``Let'' declarations *}
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   296
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   297
nonterminals letbinds letbind
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   298
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   299
constdefs
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14565
diff changeset
   300
  Let :: "['a::{}, 'a => 'b] => ('b::{})"
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   301
    "Let(s, f) == f(s)"
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   302
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   303
syntax
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   304
  "_bind"       :: "[pttrn, 'a] => letbind"           ("(2_ =/ _)" 10)
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   305
  ""            :: "letbind => letbinds"              ("_")
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   306
  "_binds"      :: "[letbind, letbinds] => letbinds"  ("_;/ _")
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   307
  "_Let"        :: "[letbinds, 'a] => 'a"             ("(let (_)/ in (_))" 10)
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   308
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   309
translations
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   310
  "_Let(_binds(b, bs), e)"  == "_Let(b, _Let(bs, e))"
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   311
  "let x = a in e"          == "Let(a, %x. e)"
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   312
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   313
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   314
lemma LetI: 
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   315
    assumes prem: "(!!x. x=t ==> P(u(x)))"
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   316
    shows "P(let x=t in u(x))"
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   317
apply (unfold Let_def)
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   318
apply (rule refl [THEN prem])
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   319
done
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   320
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   321
ML
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   322
{*
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   323
val Let_def = thm "Let_def";
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   324
val LetI = thm "LetI";
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   325
*}
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   326
4854
d1850e0964f2 tuned setup;
wenzelm
parents: 4793
diff changeset
   327
end