src/FOL/IFOL.thy
author nipkow
Fri, 11 Dec 2020 17:58:01 +0100
changeset 72884 50f18a822ee9
parent 71959 ee2c7f0dd1be
child 73015 2d7060a3ea11
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1268
f6ef556b3ede corrected title
clasohm
parents: 928
diff changeset
     1
(*  Title:      FOL/IFOL.thy
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
     2
    Author:     Lawrence C Paulson and Markus Wenzel
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
     3
*)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     4
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
     5
section \<open>Intuitionistic first-order logic\<close>
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     6
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15377
diff changeset
     7
theory IFOL
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
     8
  imports Pure
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
     9
abbrevs "?<" = "\<exists>\<^sub>\<le>\<^sub>1"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
    10
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
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    13
ML_file \<open>~~/src/Tools/misc_legacy.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    14
ML_file \<open>~~/src/Provers/splitter.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    15
ML_file \<open>~~/src/Provers/hypsubst.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    16
ML_file \<open>~~/src/Tools/IsaPlanner/zipper.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    17
ML_file \<open>~~/src/Tools/IsaPlanner/isand.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    18
ML_file \<open>~~/src/Tools/IsaPlanner/rw_inst.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    19
ML_file \<open>~~/src/Provers/quantifier1.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    20
ML_file \<open>~~/src/Tools/intuitionistic.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    21
ML_file \<open>~~/src/Tools/project_rule.ML\<close>
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
    22
ML_file \<open>~~/src/Tools/atomize_elim.ML\<close>
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46972
diff changeset
    23
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
    25
subsection \<open>Syntax and axiomatic basis\<close>
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
    26
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39159
diff changeset
    27
setup Pure_Thy.old_appl_syntax_setup
70880
de2e2382bc0d set_preproc for object-logics with type classes;
wenzelm
parents: 70432
diff changeset
    28
setup \<open>Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\<close>
26956
1309a6a0a29f setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents: 26580
diff changeset
    29
55380
4de48353034e prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents: 52241
diff changeset
    30
class "term"
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    31
default_sort \<open>term\<close>
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    33
typedecl o
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    34
11747
17a6dcd6f3cf judgment Trueprop;
wenzelm
parents: 11734
diff changeset
    35
judgment
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    36
  Trueprop :: \<open>o \<Rightarrow> prop\<close>  (\<open>(_)\<close> 5)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    38
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
    39
subsubsection \<open>Equality\<close>
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    40
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    41
axiomatization
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    42
  eq :: \<open>['a, 'a] \<Rightarrow> o\<close>  (infixl \<open>=\<close> 50)
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    43
where
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    44
  refl: \<open>a = a\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    45
  subst: \<open>a = b \<Longrightarrow> P(a) \<Longrightarrow> P(b)\<close>
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    46
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
    48
subsubsection \<open>Propositional logic\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    49
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    50
axiomatization
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    51
  False :: \<open>o\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    52
  conj :: \<open>[o, o] => o\<close>  (infixr \<open>\<and>\<close> 35) and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    53
  disj :: \<open>[o, o] => o\<close>  (infixr \<open>\<or>\<close> 30) and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    54
  imp :: \<open>[o, o] => o\<close>  (infixr \<open>\<longrightarrow>\<close> 25)
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    55
where
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    56
  conjI: \<open>\<lbrakk>P;  Q\<rbrakk> \<Longrightarrow> P \<and> Q\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    57
  conjunct1: \<open>P \<and> Q \<Longrightarrow> P\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    58
  conjunct2: \<open>P \<and> Q \<Longrightarrow> Q\<close> and
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    59
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    60
  disjI1: \<open>P \<Longrightarrow> P \<or> Q\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    61
  disjI2: \<open>Q \<Longrightarrow> P \<or> Q\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    62
  disjE: \<open>\<lbrakk>P \<or> Q; P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close> and
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    63
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    64
  impI: \<open>(P \<Longrightarrow> Q) \<Longrightarrow> P \<longrightarrow> Q\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    65
  mp: \<open>\<lbrakk>P \<longrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close> and
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    66
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    67
  FalseE: \<open>False \<Longrightarrow> P\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    68
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    69
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
    70
subsubsection \<open>Quantifiers\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    71
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    72
axiomatization
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    73
  All :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close>  (binder \<open>\<forall>\<close> 10) and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    74
  Ex :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close>  (binder \<open>\<exists>\<close> 10)
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    75
where
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    76
  allI: \<open>(\<And>x. P(x)) \<Longrightarrow> (\<forall>x. P(x))\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    77
  spec: \<open>(\<forall>x. P(x)) \<Longrightarrow> P(x)\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    78
  exI: \<open>P(x) \<Longrightarrow> (\<exists>x. P(x))\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    79
  exE: \<open>\<lbrakk>\<exists>x. P(x); \<And>x. P(x) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    80
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    81
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
    82
subsubsection \<open>Definitions\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    83
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    84
definition \<open>True \<equiv> False \<longrightarrow> False\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
    85
69587
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
    86
definition Not (\<open>\<not> _\<close> [40] 40)
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    87
  where not_def: \<open>\<not> P \<equiv> P \<longrightarrow> False\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    88
69587
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
    89
definition iff  (infixr \<open>\<longleftrightarrow>\<close> 25)
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    90
  where \<open>P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
    91
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
    92
definition Uniq :: "('a \<Rightarrow> o) \<Rightarrow> o"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
    93
  where \<open>Uniq(P) \<equiv> (\<forall>x y. P(x) \<longrightarrow> P(y) \<longrightarrow> y = x)\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
    94
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    95
definition Ex1 :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close>  (binder \<open>\<exists>!\<close> 10)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    96
  where ex1_def: \<open>\<exists>!x. P(x) \<equiv> \<exists>x. P(x) \<and> (\<forall>y. P(y) \<longrightarrow> y = x)\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
    97
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
    98
axiomatization where  \<comment> \<open>Reflection, admissible\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
    99
  eq_reflection: \<open>(x = y) \<Longrightarrow> (x \<equiv> y)\<close> and
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   100
  iff_reflection: \<open>(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   101
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   102
abbreviation not_equal :: \<open>['a, 'a] \<Rightarrow> o\<close>  (infixl \<open>\<noteq>\<close> 50)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   103
  where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close>
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
   104
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   105
syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o"  ("(2\<exists>\<^sub>\<le>\<^sub>1 _./ _)" [0, 10] 10)
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   106
translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   107
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   108
print_translation \<open>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   109
 [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\<open>Uniq\<close> \<^syntax_const>\<open>_Uniq\<close>]
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   110
\<close> \<comment> \<open>to avoid eta-contraction of body\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   111
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
   112
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   113
subsubsection \<open>Old-style ASCII syntax\<close>
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   114
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   115
notation (ASCII)
69587
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   116
  not_equal  (infixl \<open>~=\<close> 50) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   117
  Not  (\<open>~ _\<close> [40] 40) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   118
  conj  (infixr \<open>&\<close> 35) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   119
  disj  (infixr \<open>|\<close> 30) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   120
  All  (binder \<open>ALL \<close> 10) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   121
  Ex  (binder \<open>EX \<close> 10) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   122
  Ex1  (binder \<open>EX! \<close> 10) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   123
  imp  (infixr \<open>-->\<close> 25) and
53982d5ec0bb isabelle update -u mixfix_cartouches;
wenzelm
parents: 68816
diff changeset
   124
  iff  (infixr \<open><->\<close> 25)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
   125
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   126
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   127
subsection \<open>Lemmas and proof tools\<close>
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   128
46972
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
   129
lemmas strip = impI allI
ef6fc1a0884d modernized axiomatization;
wenzelm
parents: 44121
diff changeset
   130
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   131
lemma TrueI: \<open>True\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   132
  unfolding True_def by (rule impI)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   133
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   134
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   135
subsubsection \<open>Sequent-style elimination rules for \<open>\<and>\<close> \<open>\<longrightarrow>\<close> and \<open>\<forall>\<close>\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   136
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   137
lemma conjE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   138
  assumes major: \<open>P \<and> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   139
    and r: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   140
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   141
proof (rule r)
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   142
  show "P"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   143
    by (rule major [THEN conjunct1])
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   144
  show "Q"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   145
    by (rule major [THEN conjunct2])
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   146
qed
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   147
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   148
lemma impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   149
  assumes major: \<open>P \<longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   150
    and \<open>P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   151
  and r: \<open>Q \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   152
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   153
proof (rule r)
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   154
  show "Q"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   155
    by (rule mp [OF major \<open>P\<close>])
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   156
qed
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   157
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   158
lemma allE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   159
  assumes major: \<open>\<forall>x. P(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   160
    and r: \<open>P(x) \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   161
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   162
proof (rule r)
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   163
  show "P(x)"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   164
    by (rule major [THEN spec])
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   165
qed
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   166
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
   167
text \<open>Duplicates the quantifier; for use with \<^ML>\<open>eresolve_tac\<close>.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   168
lemma all_dupE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   169
  assumes major: \<open>\<forall>x. P(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   170
    and r: \<open>\<lbrakk>P(x); \<forall>x. P(x)\<rbrakk> \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   171
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   172
proof (rule r)
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   173
  show "P(x)"
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   174
    by (rule major [THEN spec])
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   175
qed (rule major)
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   176
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   177
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   178
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   179
subsubsection \<open>Negation rules, which translate between \<open>\<not> P\<close> and \<open>P \<longrightarrow> False\<close>\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   180
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   181
lemma notI: \<open>(P \<Longrightarrow> False) \<Longrightarrow> \<not> P\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   182
  unfolding not_def by (erule impI)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   183
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   184
lemma notE: \<open>\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   185
  unfolding not_def by (erule mp [THEN FalseE])
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   186
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   187
lemma rev_notE: \<open>\<lbrakk>P; \<not> P\<rbrakk> \<Longrightarrow> R\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   188
  by (erule notE)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   189
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   190
text \<open>This is useful with the special implication rules for each kind of \<open>P\<close>.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   191
lemma not_to_imp:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   192
  assumes \<open>\<not> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   193
    and r: \<open>P \<longrightarrow> False \<Longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   194
  shows \<open>Q\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   195
  apply (rule r)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   196
  apply (rule impI)
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   197
  apply (erule notE [OF \<open>\<not> P\<close>])
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   198
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   199
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   200
text \<open>
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   201
  For substitution into an assumption \<open>P\<close>, reduce \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, substitute into this implication, then apply \<open>impI\<close> to
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   202
  move \<open>P\<close> back into the assumptions.
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   203
\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   204
lemma rev_mp: \<open>\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   205
  by (erule mp)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   206
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   207
text \<open>Contrapositive of an inference rule.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   208
lemma contrapos:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   209
  assumes major: \<open>\<not> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   210
    and minor: \<open>P \<Longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   211
  shows \<open>\<not> P\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   212
  apply (rule major [THEN notE, THEN notI])
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   213
  apply (erule minor)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   214
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   215
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   216
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   217
subsubsection \<open>Modus Ponens Tactics\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   218
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   219
text \<open>
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   220
  Finds \<open>P \<longrightarrow> Q\<close> and P in the assumptions, replaces implication by
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   221
  \<open>Q\<close>.
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   222
\<close>
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   223
ML \<open>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   224
  fun mp_tac ctxt i =
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   225
    eresolve_tac ctxt @{thms notE impE} i THEN assume_tac ctxt i;
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   226
  fun eq_mp_tac ctxt i =
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   227
    eresolve_tac ctxt @{thms notE impE} i THEN eq_assume_tac i;
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   228
\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   229
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   230
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   231
subsection \<open>If-and-only-if\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   232
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   233
lemma iffI: \<open>\<lbrakk>P \<Longrightarrow> Q; Q \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> Q\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   234
  unfolding iff_def
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   235
  by (rule conjI; erule impI)
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   236
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   237
lemma iffE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   238
  assumes major: \<open>P \<longleftrightarrow> Q\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   239
    and r: \<open>\<lbrakk>P \<longrightarrow> Q; Q \<longrightarrow> P\<rbrakk> \<Longrightarrow> R\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   240
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   241
  using major
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   242
  unfolding iff_def
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   243
  apply (rule conjE)
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   244
  apply (erule r)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   245
  apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   246
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   247
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   248
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   249
subsubsection \<open>Destruct rules for \<open>\<longleftrightarrow>\<close> similar to Modus Ponens\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   250
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   251
lemma iffD1: \<open>\<lbrakk>P \<longleftrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   252
  unfolding iff_def
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   253
  apply (erule conjunct1 [THEN mp])
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   254
  apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   255
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   256
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   257
lemma iffD2: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q\<rbrakk> \<Longrightarrow> P\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   258
  unfolding iff_def
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   259
  apply (erule conjunct2 [THEN mp])
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   260
  apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   261
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   262
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   263
lemma rev_iffD1: \<open>\<lbrakk>P; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   264
  apply (erule iffD1)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   265
  apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   266
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   267
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   268
lemma rev_iffD2: \<open>\<lbrakk>Q; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> P\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   269
  apply (erule iffD2)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   270
  apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   271
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   272
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   273
lemma iff_refl: \<open>P \<longleftrightarrow> P\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   274
  by (rule iffI)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   275
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   276
lemma iff_sym: \<open>Q \<longleftrightarrow> P \<Longrightarrow> P \<longleftrightarrow> Q\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   277
  apply (erule iffE)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   278
  apply (rule iffI)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   279
  apply (assumption | erule mp)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   280
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   281
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   282
lemma iff_trans: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q \<longleftrightarrow> R\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> R\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   283
  apply (rule iffI)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   284
  apply (assumption | erule iffE | erule (1) notE impE)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   285
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   286
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   287
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   288
subsection \<open>Unique existence\<close>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   289
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   290
text \<open>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   291
  NOTE THAT the following 2 quantifications:
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   292
63906
fa799a8e4adc repair LaTeX dropout from f83ef97d8d7d
Lars Hupel <lars.hupel@mytum.de>
parents: 63901
diff changeset
   293
    \<^item> \<open>\<exists>!x\<close> such that [\<open>\<exists>!y\<close> such that P(x,y)]   (sequential)
fa799a8e4adc repair LaTeX dropout from f83ef97d8d7d
Lars Hupel <lars.hupel@mytum.de>
parents: 63901
diff changeset
   294
    \<^item> \<open>\<exists>!x,y\<close> such that P(x,y)                   (simultaneous)
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   295
63906
fa799a8e4adc repair LaTeX dropout from f83ef97d8d7d
Lars Hupel <lars.hupel@mytum.de>
parents: 63901
diff changeset
   296
  do NOT mean the same thing. The parser treats \<open>\<exists>!x y.P(x,y)\<close> as sequential.
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   297
\<close>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   298
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   299
lemma ex1I: \<open>P(a) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> x = a) \<Longrightarrow> \<exists>!x. P(x)\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   300
  unfolding ex1_def
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   301
  apply (assumption | rule exI conjI allI impI)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   302
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   303
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   304
text \<open>Sometimes easier to use: the premises have no shared variables. Safe!\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   305
lemma ex_ex1I: \<open>\<exists>x. P(x) \<Longrightarrow> (\<And>x y. \<lbrakk>P(x); P(y)\<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> \<exists>!x. P(x)\<close>
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   306
  apply (erule exE)
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   307
  apply (rule ex1I)
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   308
   apply assumption
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   309
  apply assumption
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   310
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   311
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   312
lemma ex1E: \<open>\<exists>! x. P(x) \<Longrightarrow> (\<And>x. \<lbrakk>P(x); \<forall>y. P(y) \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   313
  unfolding ex1_def
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   314
  apply (assumption | erule exE conjE)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   315
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   316
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   317
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   318
subsubsection \<open>\<open>\<longleftrightarrow>\<close> congruence rules for simplification\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   319
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   320
text \<open>Use \<open>iffE\<close> on a premise. For \<open>conj_cong\<close>, \<open>imp_cong\<close>, \<open>all_cong\<close>, \<open>ex_cong\<close>.\<close>
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   321
ML \<open>
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   322
  fun iff_tac ctxt prems i =
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   323
    resolve_tac ctxt (prems RL @{thms iffE}) i THEN
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   324
    REPEAT1 (eresolve_tac ctxt @{thms asm_rl mp} i);
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   325
\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   326
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   327
method_setup iff =
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   328
  \<open>Attrib.thms >>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   329
    (fn prems => fn ctxt => SIMPLE_METHOD' (iff_tac ctxt prems))\<close>
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   330
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   331
lemma conj_cong:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   332
  assumes \<open>P \<longleftrightarrow> P'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   333
    and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   334
  shows \<open>(P \<and> Q) \<longleftrightarrow> (P' \<and> Q')\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   335
  apply (insert assms)
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   336
  apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   337
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   338
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   339
text \<open>Reversed congruence rule!  Used in ZF/Order.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   340
lemma conj_cong2:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   341
  assumes \<open>P \<longleftrightarrow> P'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   342
    and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   343
  shows \<open>(Q \<and> P) \<longleftrightarrow> (Q' \<and> P')\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   344
  apply (insert assms)
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   345
  apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   346
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   347
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   348
lemma disj_cong:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   349
  assumes \<open>P \<longleftrightarrow> P'\<close> and \<open>Q \<longleftrightarrow> Q'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   350
  shows \<open>(P \<or> Q) \<longleftrightarrow> (P' \<or> Q')\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   351
  apply (insert assms)
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   352
  apply (erule iffE disjE disjI1 disjI2 |
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   353
    assumption | rule iffI | erule (1) notE impE)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   354
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   355
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   356
lemma imp_cong:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   357
  assumes \<open>P \<longleftrightarrow> P'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   358
    and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   359
  shows \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (P' \<longrightarrow> Q')\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   360
  apply (insert assms)
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   361
  apply (assumption | rule iffI impI | erule iffE | erule (1) notE impE | iff assms)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   362
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   363
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   364
lemma iff_cong: \<open>\<lbrakk>P \<longleftrightarrow> P'; Q \<longleftrightarrow> Q'\<rbrakk> \<Longrightarrow> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P' \<longleftrightarrow> Q')\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   365
  apply (erule iffE | assumption | rule iffI | erule (1) notE impE)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   366
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   367
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   368
lemma not_cong: \<open>P \<longleftrightarrow> P' \<Longrightarrow> \<not> P \<longleftrightarrow> \<not> P'\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   369
  apply (assumption | rule iffI notI | erule (1) notE impE | erule iffE notE)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   370
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   371
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   372
lemma all_cong:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   373
  assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   374
  shows \<open>(\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))\<close>
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   375
  apply (assumption | rule iffI allI | erule (1) notE impE | erule allE | iff assms)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   376
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   377
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   378
lemma ex_cong:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   379
  assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   380
  shows \<open>(\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))\<close>
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   381
  apply (erule exE | assumption | rule iffI exI | erule (1) notE impE | iff assms)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   382
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   383
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   384
lemma ex1_cong:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   385
  assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   386
  shows \<open>(\<exists>!x. P(x)) \<longleftrightarrow> (\<exists>!x. Q(x))\<close>
59529
d881f5288d5a proper context and method setup;
wenzelm
parents: 59498
diff changeset
   387
  apply (erule ex1E spec [THEN mp] | assumption | rule iffI ex1I | erule (1) notE impE | iff assms)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   388
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   389
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   390
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   391
subsection \<open>Equality rules\<close>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   392
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   393
lemma sym: \<open>a = b \<Longrightarrow> b = a\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   394
  apply (erule subst)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   395
  apply (rule refl)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   396
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   397
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   398
lemma trans: \<open>\<lbrakk>a = b; b = c\<rbrakk> \<Longrightarrow> a = c\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   399
  apply (erule subst, assumption)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   400
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   401
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   402
lemma not_sym: \<open>b \<noteq> a \<Longrightarrow> a \<noteq> b\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   403
  apply (erule contrapos)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   404
  apply (erule sym)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   405
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   406
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   407
text \<open>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   408
  Two theorems for rewriting only one instance of a definition:
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   409
  the first for definitions of formulae and the second for terms.
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   410
\<close>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   411
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   412
lemma def_imp_iff: \<open>(A \<equiv> B) \<Longrightarrow> A \<longleftrightarrow> B\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   413
  apply unfold
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   414
  apply (rule iff_refl)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   415
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   416
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   417
lemma meta_eq_to_obj_eq: \<open>(A \<equiv> B) \<Longrightarrow> A = B\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   418
  apply unfold
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   419
  apply (rule refl)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   420
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   421
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   422
lemma meta_eq_to_iff: \<open>x \<equiv> y \<Longrightarrow> x \<longleftrightarrow> y\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   423
  by unfold (rule iff_refl)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   424
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   425
text \<open>Substitution.\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   426
lemma ssubst: \<open>\<lbrakk>b = a; P(a)\<rbrakk> \<Longrightarrow> P(b)\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   427
  apply (drule sym)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   428
  apply (erule (1) subst)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   429
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   430
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   431
text \<open>A special case of \<open>ex1E\<close> that would otherwise need quantifier
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   432
  expansion.\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   433
lemma ex1_equalsE: \<open>\<lbrakk>\<exists>!x. P(x); P(a); P(b)\<rbrakk> \<Longrightarrow> a = b\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   434
  apply (erule ex1E)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   435
  apply (rule trans)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   436
   apply (rule_tac [2] sym)
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   437
   apply (assumption | erule spec [THEN mp])+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   438
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   439
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   440
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   441
subsection \<open>Simplifications of assumed implications\<close>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   442
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   443
text \<open>
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   444
  Roy Dyckhoff has proved that \<open>conj_impE\<close>, \<open>disj_impE\<close>, and
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
   445
  \<open>imp_impE\<close> used with \<^ML>\<open>mp_tac\<close> (restricted to atomic formulae) is
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   446
  COMPLETE for intuitionistic propositional logic.
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   447
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   448
  See R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   449
  (preprint, University of St Andrews, 1991).
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   450
\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   451
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   452
lemma conj_impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   453
  assumes major: \<open>(P \<and> Q) \<longrightarrow> S\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   454
    and r: \<open>P \<longrightarrow> (Q \<longrightarrow> S) \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   455
  shows \<open>R\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   456
  by (assumption | rule conjI impI major [THEN mp] r)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   457
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   458
lemma disj_impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   459
  assumes major: \<open>(P \<or> Q) \<longrightarrow> S\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   460
    and r: \<open>\<lbrakk>P \<longrightarrow> S; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   461
  shows \<open>R\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   462
  by (assumption | rule disjI1 disjI2 impI major [THEN mp] r)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   463
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   464
text \<open>Simplifies the implication.  Classical version is stronger.
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   465
  Still UNSAFE since Q must be provable -- backtracking needed.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   466
lemma imp_impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   467
  assumes major: \<open>(P \<longrightarrow> Q) \<longrightarrow> S\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   468
    and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   469
    and r2: \<open>S \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   470
  shows \<open>R\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   471
  by (assumption | rule impI major [THEN mp] r1 r2)+
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   472
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   473
text \<open>Simplifies the implication.  Classical version is stronger.
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   474
  Still UNSAFE since ~P must be provable -- backtracking needed.\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   475
lemma not_impE: \<open>\<not> P \<longrightarrow> S \<Longrightarrow> (P \<Longrightarrow> False) \<Longrightarrow> (S \<Longrightarrow> R) \<Longrightarrow> R\<close>
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   476
  apply (drule mp)
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   477
   apply (rule notI | assumption)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   478
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   479
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   480
text \<open>Simplifies the implication. UNSAFE.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   481
lemma iff_impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   482
  assumes major: \<open>(P \<longleftrightarrow> Q) \<longrightarrow> S\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   483
    and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   484
    and r2: \<open>\<lbrakk>Q; P \<longrightarrow> S\<rbrakk> \<Longrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   485
    and r3: \<open>S \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   486
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   487
  by (assumption | rule iffI impI major [THEN mp] r1 r2 r3)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   488
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   489
text \<open>What if \<open>(\<forall>x. \<not> \<not> P(x)) \<longrightarrow> \<not> \<not> (\<forall>x. P(x))\<close> is an assumption?
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   490
  UNSAFE.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   491
lemma all_impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   492
  assumes major: \<open>(\<forall>x. P(x)) \<longrightarrow> S\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   493
    and r1: \<open>\<And>x. P(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   494
    and r2: \<open>S \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   495
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   496
  by (rule allI impI major [THEN mp] r1 r2)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   497
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   498
text \<open>
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   499
  Unsafe: \<open>\<exists>x. P(x)) \<longrightarrow> S\<close> is equivalent
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   500
  to \<open>\<forall>x. P(x) \<longrightarrow> S\<close>.\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   501
lemma ex_impE:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   502
  assumes major: \<open>(\<exists>x. P(x)) \<longrightarrow> S\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   503
    and r: \<open>P(x) \<longrightarrow> S \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   504
  shows \<open>R\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   505
  by (assumption | rule exI impI major [THEN mp] r)+
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   506
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   507
text \<open>Courtesy of Krzysztof Grabczewski.\<close>
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   508
lemma disj_imp_disj: \<open>P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> S) \<Longrightarrow> R \<or> S\<close>
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23171
diff changeset
   509
  apply (erule disjE)
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   510
  apply (rule disjI1) apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   511
  apply (rule disjI2) apply assumption
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   512
  done
11734
7a21bf539412 declare impE iffD1 iffD2 ad elim of Pure;
wenzelm
parents: 11677
diff changeset
   513
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   514
ML \<open>
32172
c4e55f30d527 renamed functor ProjectRuleFun to Project_Rule;
wenzelm
parents: 31299
diff changeset
   515
structure Project_Rule = Project_Rule
c4e55f30d527 renamed functor ProjectRuleFun to Project_Rule;
wenzelm
parents: 31299
diff changeset
   516
(
22139
539a63b98f76 tuned ML setup;
wenzelm
parents: 21539
diff changeset
   517
  val conjunct1 = @{thm conjunct1}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21539
diff changeset
   518
  val conjunct2 = @{thm conjunct2}
539a63b98f76 tuned ML setup;
wenzelm
parents: 21539
diff changeset
   519
  val mp = @{thm mp}
32172
c4e55f30d527 renamed functor ProjectRuleFun to Project_Rule;
wenzelm
parents: 31299
diff changeset
   520
)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   521
\<close>
18481
b75ce99617c7 structure ProjectRule;
wenzelm
parents: 17702
diff changeset
   522
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   523
ML_file \<open>fologic.ML\<close>
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   524
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   525
lemma thin_refl: \<open>\<lbrakk>x = x; PROP W\<rbrakk> \<Longrightarrow> PROP W\<close> .
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   526
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   527
ML \<open>
42799
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   528
structure Hypsubst = Hypsubst
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   529
(
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   530
  val dest_eq = FOLogic.dest_eq
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   531
  val dest_Trueprop = FOLogic.dest_Trueprop
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   532
  val dest_imp = FOLogic.dest_imp
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   533
  val eq_reflection = @{thm eq_reflection}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   534
  val rev_eq_reflection = @{thm meta_eq_to_obj_eq}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   535
  val imp_intr = @{thm impI}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   536
  val rev_mp = @{thm rev_mp}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   537
  val subst = @{thm subst}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   538
  val sym = @{thm sym}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   539
  val thin_refl = @{thm thin_refl}
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   540
);
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   541
open Hypsubst;
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   542
\<close>
42799
4e33894aec6d modernized functor names;
wenzelm
parents: 42303
diff changeset
   543
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69593
diff changeset
   544
ML_file \<open>intprover.ML\<close>
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   545
4092
9faf228771dc added simpset thy_data;
wenzelm
parents: 3932
diff changeset
   546
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   547
subsection \<open>Intuitionistic Reasoning\<close>
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   548
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69590
diff changeset
   549
setup \<open>Intuitionistic.method_setup \<^binding>\<open>iprover\<close>\<close>
30165
6ee87f67d9cd moved generic intuitionistic prover to src/Tools/intuitionistic.ML;
wenzelm
parents: 30160
diff changeset
   550
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   551
lemma impE':
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   552
  assumes 1: \<open>P \<longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   553
    and 2: \<open>Q \<Longrightarrow> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   554
    and 3: \<open>P \<longrightarrow> Q \<Longrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   555
  shows \<open>R\<close>
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   556
proof -
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   557
  from 3 and 1 have \<open>P\<close> .
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   558
  with 1 have \<open>Q\<close> by (rule impE)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   559
  with 2 show \<open>R\<close> .
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   560
qed
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   561
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   562
lemma allE':
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   563
  assumes 1: \<open>\<forall>x. P(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   564
    and 2: \<open>P(x) \<Longrightarrow> \<forall>x. P(x) \<Longrightarrow> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   565
  shows \<open>Q\<close>
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   566
proof -
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   567
  from 1 have \<open>P(x)\<close> by (rule spec)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   568
  from this and 1 show \<open>Q\<close> by (rule 2)
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   569
qed
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   570
12937
0c4fd7529467 clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents: 12875
diff changeset
   571
lemma notE':
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   572
  assumes 1: \<open>\<not> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   573
    and 2: \<open>\<not> P \<Longrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   574
  shows \<open>R\<close>
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   575
proof -
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   576
  from 2 and 1 have \<open>P\<close> .
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   577
  with 1 show \<open>R\<close> by (rule notE)
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   578
qed
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   579
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   580
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   581
  and [Pure.intro!] = iffI conjI impI TrueI notI allI refl
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   582
  and [Pure.elim 2] = allE notE' impE'
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   583
  and [Pure.intro] = exI disjI2 disjI1
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   584
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   585
setup \<open>
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   586
  Context_Rules.addSWrapper
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   587
    (fn ctxt => fn tac => hyp_subst_tac ctxt ORELSE' tac)
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   588
\<close>
12349
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   589
94e812f9683e setup "rules" method;
wenzelm
parents: 12114
diff changeset
   590
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   591
lemma iff_not_sym: \<open>\<not> (Q \<longleftrightarrow> P) \<Longrightarrow> \<not> (P \<longleftrightarrow> Q)\<close>
17591
33d409318266 renamed rules to iprover
nipkow
parents: 17276
diff changeset
   592
  by iprover
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   593
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   594
lemmas [sym] = sym iff_sym not_sym iff_not_sym
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   595
  and [Pure.elim?] = iffD1 iffD2 impE
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   596
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   597
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   598
lemma eq_commute: \<open>a = b \<longleftrightarrow> b = a\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   599
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   600
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   601
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   602
subsection \<open>Polymorphic congruence rules\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   603
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   604
lemma subst_context: \<open>a = b \<Longrightarrow> t(a) = t(b)\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   605
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   606
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   607
lemma subst_context2: \<open>\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> t(a,c) = t(b,d)\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   608
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   609
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   610
lemma subst_context3: \<open>\<lbrakk>a = b; c = d; e = f\<rbrakk> \<Longrightarrow> t(a,c,e) = t(b,d,f)\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   611
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   612
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   613
text \<open>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   614
  Useful with \<^ML>\<open>eresolve_tac\<close> for proving equalities from known
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   615
  equalities.
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   616
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   617
        a = b
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   618
        |   |
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   619
        c = d
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   620
\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   621
lemma box_equals: \<open>\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   622
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   623
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   624
text \<open>Dual of \<open>box_equals\<close>: for proving equalities backwards.\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   625
lemma simp_equals: \<open>\<lbrakk>a = c; b = d; c = d\<rbrakk> \<Longrightarrow> a = b\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   626
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   627
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   628
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   629
subsubsection \<open>Congruence rules for predicate letters\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   630
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   631
lemma pred1_cong: \<open>a = a' \<Longrightarrow> P(a) \<longleftrightarrow> P(a')\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   632
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   633
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   634
lemma pred2_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> P(a,b) \<longleftrightarrow> P(a',b')\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   635
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   636
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   637
lemma pred3_cong: \<open>\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> P(a,b,c) \<longleftrightarrow> P(a',b',c')\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   638
  by iprover
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   639
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   640
text \<open>Special case for the equality predicate!\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   641
lemma eq_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> a = b \<longleftrightarrow> a' = b'\<close>
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   642
  by iprover
13435
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   643
05631e8f0258 new theorem eq_commute
paulson
parents: 12937
diff changeset
   644
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   645
subsection \<open>Atomizing meta-level rules\<close>
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   646
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   647
lemma atomize_all [atomize]: \<open>(\<And>x. P(x)) \<equiv> Trueprop (\<forall>x. P(x))\<close>
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   648
proof
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   649
  assume \<open>\<And>x. P(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   650
  then show \<open>\<forall>x. P(x)\<close> ..
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   651
next
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   652
  assume \<open>\<forall>x. P(x)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   653
  then show \<open>\<And>x. P(x)\<close> ..
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   654
qed
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   655
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   656
lemma atomize_imp [atomize]: \<open>(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)\<close>
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   657
proof
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   658
  assume \<open>A \<Longrightarrow> B\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   659
  then show \<open>A \<longrightarrow> B\<close> ..
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   660
next
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   661
  assume \<open>A \<longrightarrow> B\<close> and \<open>A\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   662
  then show \<open>B\<close> by (rule mp)
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   663
qed
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   664
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   665
lemma atomize_eq [atomize]: \<open>(x \<equiv> y) \<equiv> Trueprop (x = y)\<close>
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   666
proof
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   667
  assume \<open>x \<equiv> y\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   668
  show \<open>x = y\<close> unfolding \<open>x \<equiv> y\<close> by (rule refl)
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   669
next
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   670
  assume \<open>x = y\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   671
  then show \<open>x \<equiv> y\<close> by (rule eq_reflection)
11677
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   672
qed
ee12f18599e5 atomize stuff from theory FOL;
wenzelm
parents: 9886
diff changeset
   673
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   674
lemma atomize_iff [atomize]: \<open>(A \<equiv> B) \<equiv> Trueprop (A \<longleftrightarrow> B)\<close>
18813
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   675
proof
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   676
  assume \<open>A \<equiv> B\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   677
  show \<open>A \<longleftrightarrow> B\<close> unfolding \<open>A \<equiv> B\<close> by (rule iff_refl)
18813
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   678
next
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   679
  assume \<open>A \<longleftrightarrow> B\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   680
  then show \<open>A \<equiv> B\<close> by (rule iff_reflection)
18813
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   681
qed
fc35dcc2558b added atomize_iff;
wenzelm
parents: 18708
diff changeset
   682
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   683
lemma atomize_conj [atomize]: \<open>(A &&& B) \<equiv> Trueprop (A \<and> B)\<close>
11976
075df6e46cef equal_intr_rule already declared in Pure;
wenzelm
parents: 11953
diff changeset
   684
proof
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   685
  assume conj: \<open>A &&& B\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   686
  show \<open>A \<and> B\<close>
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   687
  proof (rule conjI)
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   688
    from conj show \<open>A\<close> by (rule conjunctionD1)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   689
    from conj show \<open>B\<close> by (rule conjunctionD2)
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   690
  qed
11953
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   691
next
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   692
  assume conj: \<open>A \<and> B\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   693
  show \<open>A &&& B\<close>
19120
353d349740de not_equal: replaced syntax translation by abbreviation;
wenzelm
parents: 18861
diff changeset
   694
  proof -
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   695
    from conj show \<open>A\<close> ..
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   696
    from conj show \<open>B\<close> ..
11953
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   697
  qed
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   698
qed
f98623fdf6ef atomize_conj;
wenzelm
parents: 11848
diff changeset
   699
12368
2af9ad81ea56 sym declarations;
wenzelm
parents: 12352
diff changeset
   700
lemmas [symmetric, rulify] = atomize_all atomize_imp
18861
38269ef5175a declare defn rules;
wenzelm
parents: 18813
diff changeset
   701
  and [symmetric, defn] = atomize_all atomize_imp atomize_eq atomize_iff
11771
b7b100a2de1d moved rulify to ObjectLogic;
wenzelm
parents: 11747
diff changeset
   702
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   703
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   704
subsection \<open>Atomizing elimination rules\<close>
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 26286
diff changeset
   705
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   706
lemma atomize_exL[atomize_elim]: \<open>(\<And>x. P(x) \<Longrightarrow> Q) \<equiv> ((\<exists>x. P(x)) \<Longrightarrow> Q)\<close>
57948
75724d71013c modernized module name and setup;
wenzelm
parents: 55380
diff changeset
   707
  by rule iprover+
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 26286
diff changeset
   708
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   709
lemma atomize_conjL[atomize_elim]: \<open>(A \<Longrightarrow> B \<Longrightarrow> C) \<equiv> (A \<and> B \<Longrightarrow> C)\<close>
57948
75724d71013c modernized module name and setup;
wenzelm
parents: 55380
diff changeset
   710
  by rule iprover+
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 26286
diff changeset
   711
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   712
lemma atomize_disjL[atomize_elim]: \<open>((A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C) \<equiv> ((A \<or> B \<Longrightarrow> C) \<Longrightarrow> C)\<close>
57948
75724d71013c modernized module name and setup;
wenzelm
parents: 55380
diff changeset
   713
  by rule iprover+
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 26286
diff changeset
   714
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   715
lemma atomize_elimL[atomize_elim]: \<open>(\<And>B. (A \<Longrightarrow> B) \<Longrightarrow> B) \<equiv> Trueprop(A)\<close> ..
26580
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 26286
diff changeset
   716
c3e597a476fd Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents: 26286
diff changeset
   717
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   718
subsection \<open>Calculational rules\<close>
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   719
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   720
lemma forw_subst: \<open>a = b \<Longrightarrow> P(b) \<Longrightarrow> P(a)\<close>
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   721
  by (rule ssubst)
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   722
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   723
lemma back_subst: \<open>P(a) \<Longrightarrow> a = b \<Longrightarrow> P(b)\<close>
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   724
  by (rule subst)
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   725
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   726
text \<open>
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   727
  Note that this list of rules is in reverse order of priorities.
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   728
\<close>
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   729
12019
abe9b7c6016e transitive declared in Pure;
wenzelm
parents: 11976
diff changeset
   730
lemmas basic_trans_rules [trans] =
11848
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   731
  forw_subst
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   732
  back_subst
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   733
  rev_mp
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   734
  mp
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   735
  trans
6e3017adb8c0 calculational rules moved from FOL to IFOL;
wenzelm
parents: 11771
diff changeset
   736
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   737
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   738
subsection \<open>``Let'' declarations\<close>
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   739
41229
d797baa3d57c replaced command 'nonterminals' by slightly modernized version 'nonterminal';
wenzelm
parents: 39557
diff changeset
   740
nonterminal letbinds and letbind
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   741
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   742
definition Let :: \<open>['a::{}, 'a => 'b] \<Rightarrow> ('b::{})\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   743
  where \<open>Let(s, f) \<equiv> f(s)\<close>
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   744
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   745
syntax
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   746
  "_bind"       :: \<open>[pttrn, 'a] => letbind\<close>           (\<open>(2_ =/ _)\<close> 10)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   747
  ""            :: \<open>letbind => letbinds\<close>              (\<open>_\<close>)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   748
  "_binds"      :: \<open>[letbind, letbinds] => letbinds\<close>  (\<open>_;/ _\<close>)
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   749
  "_Let"        :: \<open>[letbinds, 'a] => 'a\<close>             (\<open>(let (_)/ in (_))\<close> 10)
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   750
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   751
translations
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   752
  "_Let(_binds(b, bs), e)"  == "_Let(b, _Let(bs, e))"
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   753
  "let x = a in e"          == "CONST Let(a, \<lambda>x. e)"
13779
2a34dc5cf79e moving "let" from ZF to FOL
paulson
parents: 13435
diff changeset
   754
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   755
lemma LetI:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   756
  assumes \<open>\<And>x. x = t \<Longrightarrow> P(u(x))\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   757
  shows \<open>P(let x = t in u(x))\<close>
71839
0bbe0866b7e6 The Uniq quantifier for FOL too
paulson <lp15@cam.ac.uk>
parents: 70880
diff changeset
   758
  unfolding Let_def
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   759
  apply (rule refl [THEN assms])
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   760
  done
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   761
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 21524
diff changeset
   762
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59529
diff changeset
   763
subsection \<open>Intuitionistic simplification rules\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   764
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   765
lemma conj_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   766
  \<open>P \<and> True \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   767
  \<open>True \<and> P \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   768
  \<open>P \<and> False \<longleftrightarrow> False\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   769
  \<open>False \<and> P \<longleftrightarrow> False\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   770
  \<open>P \<and> P \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   771
  \<open>P \<and> P \<and> Q \<longleftrightarrow> P \<and> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   772
  \<open>P \<and> \<not> P \<longleftrightarrow> False\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   773
  \<open>\<not> P \<and> P \<longleftrightarrow> False\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   774
  \<open>(P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   775
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   776
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   777
lemma disj_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   778
  \<open>P \<or> True \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   779
  \<open>True \<or> P \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   780
  \<open>P \<or> False \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   781
  \<open>False \<or> P \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   782
  \<open>P \<or> P \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   783
  \<open>P \<or> P \<or> Q \<longleftrightarrow> P \<or> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   784
  \<open>(P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   785
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   786
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   787
lemma not_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   788
  \<open>\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   789
  \<open>\<not> False \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   790
  \<open>\<not> True \<longleftrightarrow> False\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   791
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   792
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   793
lemma imp_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   794
  \<open>(P \<longrightarrow> False) \<longleftrightarrow> \<not> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   795
  \<open>(P \<longrightarrow> True) \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   796
  \<open>(False \<longrightarrow> P) \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   797
  \<open>(True \<longrightarrow> P) \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   798
  \<open>(P \<longrightarrow> P) \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   799
  \<open>(P \<longrightarrow> \<not> P) \<longleftrightarrow> \<not> P\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   800
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   801
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   802
lemma iff_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   803
  \<open>(True \<longleftrightarrow> P) \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   804
  \<open>(P \<longleftrightarrow> True) \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   805
  \<open>(P \<longleftrightarrow> P) \<longleftrightarrow> True\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   806
  \<open>(False \<longleftrightarrow> P) \<longleftrightarrow> \<not> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   807
  \<open>(P \<longleftrightarrow> False) \<longleftrightarrow> \<not> P\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   808
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   809
62020
5d208fd2507d isabelle update_cartouches -c -t;
wenzelm
parents: 61490
diff changeset
   810
text \<open>The \<open>x = t\<close> versions are needed for the simplification
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   811
  procedures.\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   812
lemma quant_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   813
  \<open>\<And>P. (\<forall>x. P) \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   814
  \<open>(\<forall>x. x = t \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   815
  \<open>(\<forall>x. t = x \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   816
  \<open>\<And>P. (\<exists>x. P) \<longleftrightarrow> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   817
  \<open>\<exists>x. x = t\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   818
  \<open>\<exists>x. t = x\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   819
  \<open>(\<exists>x. x = t \<and> P(x)) \<longleftrightarrow> P(t)\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   820
  \<open>(\<exists>x. t = x \<and> P(x)) \<longleftrightarrow> P(t)\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   821
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   822
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   823
text \<open>These are NOT supplied by default!\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   824
lemma distrib_simps:
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   825
  \<open>P \<and> (Q \<or> R) \<longleftrightarrow> P \<and> Q \<or> P \<and> R\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   826
  \<open>(Q \<or> R) \<and> P \<longleftrightarrow> Q \<and> P \<or> R \<and> P\<close>
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   827
  \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   828
  by iprover+
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   829
71919
2e7df6774373 more rules for FOL also
haftmann
parents: 71839
diff changeset
   830
lemma subst_all:
2e7df6774373 more rules for FOL also
haftmann
parents: 71839
diff changeset
   831
  \<open>(\<And>x. x = a \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close>
2e7df6774373 more rules for FOL also
haftmann
parents: 71839
diff changeset
   832
  \<open>(\<And>x. a = x \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close>
71959
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   833
proof -
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   834
  show \<open>(\<And>x. x = a \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   835
  proof (rule equal_intr_rule)
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   836
    assume *: \<open>\<And>x. x = a \<Longrightarrow> PROP P(x)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   837
    show \<open>PROP P(a)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   838
      by (rule *) (rule refl)
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   839
  next
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   840
    fix x
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   841
    assume \<open>PROP P(a)\<close> and \<open>x = a\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   842
    from \<open>x = a\<close> have \<open>x \<equiv> a\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   843
      by (rule eq_reflection)
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   844
    with \<open>PROP P(a)\<close> show \<open>PROP P(x)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   845
      by simp
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   846
  qed
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   847
  show \<open>(\<And>x. a = x \<Longrightarrow> PROP P(x)) \<equiv> PROP P(a)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   848
  proof (rule equal_intr_rule)
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   849
    assume *: \<open>\<And>x. a = x \<Longrightarrow> PROP P(x)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   850
    show \<open>PROP P(a)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   851
      by (rule *) (rule refl)
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   852
  next
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   853
    fix x
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   854
    assume \<open>PROP P(a)\<close> and \<open>a = x\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   855
    from \<open>a = x\<close> have \<open>a \<equiv> x\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   856
      by (rule eq_reflection)
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   857
    with \<open>PROP P(a)\<close> show \<open>PROP P(x)\<close>
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   858
      by simp
ee2c7f0dd1be prefer single name
haftmann
parents: 71919
diff changeset
   859
  qed
71919
2e7df6774373 more rules for FOL also
haftmann
parents: 71839
diff changeset
   860
qed
2e7df6774373 more rules for FOL also
haftmann
parents: 71839
diff changeset
   861
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   862
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   863
subsubsection \<open>Conversion into rewrite rules\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   864
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   865
lemma P_iff_F: \<open>\<not> P \<Longrightarrow> (P \<longleftrightarrow> False)\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   866
  by iprover
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   867
lemma iff_reflection_F: \<open>\<not> P \<Longrightarrow> (P \<equiv> False)\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   868
  by (rule P_iff_F [THEN iff_reflection])
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   869
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   870
lemma P_iff_T: \<open>P \<Longrightarrow> (P \<longleftrightarrow> True)\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   871
  by iprover
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   872
lemma iff_reflection_T: \<open>P \<Longrightarrow> (P \<equiv> True)\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   873
  by (rule P_iff_T [THEN iff_reflection])
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   874
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   875
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   876
subsubsection \<open>More rewrite rules\<close>
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   877
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   878
lemma conj_commute: \<open>P \<and> Q \<longleftrightarrow> Q \<and> P\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   879
lemma conj_left_commute: \<open>P \<and> (Q \<and> R) \<longleftrightarrow> Q \<and> (P \<and> R)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   880
lemmas conj_comms = conj_commute conj_left_commute
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   881
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   882
lemma disj_commute: \<open>P \<or> Q \<longleftrightarrow> Q \<or> P\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   883
lemma disj_left_commute: \<open>P \<or> (Q \<or> R) \<longleftrightarrow> Q \<or> (P \<or> R)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   884
lemmas disj_comms = disj_commute disj_left_commute
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   885
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   886
lemma conj_disj_distribL: \<open>P \<and> (Q \<or> R) \<longleftrightarrow> (P \<and> Q \<or> P \<and> R)\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   887
lemma conj_disj_distribR: \<open>(P \<or> Q) \<and> R \<longleftrightarrow> (P \<and> R \<or> Q \<and> R)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   888
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   889
lemma disj_conj_distribL: \<open>P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   890
lemma disj_conj_distribR: \<open>(P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   891
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   892
lemma imp_conj_distrib: \<open>(P \<longrightarrow> (Q \<and> R)) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   893
lemma imp_conj: \<open>((P \<and> Q) \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   894
lemma imp_disj: \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   895
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   896
lemma de_Morgan_disj: \<open>(\<not> (P \<or> Q)) \<longleftrightarrow> (\<not> P \<and> \<not> Q)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   897
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   898
lemma not_ex: \<open>(\<not> (\<exists>x. P(x))) \<longleftrightarrow> (\<forall>x. \<not> P(x))\<close> by iprover
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   899
lemma imp_ex: \<open>((\<exists>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> Q)\<close> by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   900
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   901
lemma ex_disj_distrib: \<open>(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> ((\<exists>x. P(x)) \<or> (\<exists>x. Q(x)))\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   902
  by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   903
69590
e65314985426 isabelle update_inner_syntax_cartouches;
wenzelm
parents: 69587
diff changeset
   904
lemma all_conj_distrib: \<open>(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x. P(x)) \<and> (\<forall>x. Q(x)))\<close>
61487
f8cb97e0fd0b more symbols, with swapped defaults: old-style ASCII syntax uses "ASCII" print mode;
wenzelm
parents: 61378
diff changeset
   905
  by iprover
26286
3ff5d257f175 added lemmas from simpdata.ML;
wenzelm
parents: 24830
diff changeset
   906
4854
d1850e0964f2 tuned setup;
wenzelm
parents: 4793
diff changeset
   907
end