author | wenzelm |
Thu, 12 Jun 2025 12:59:17 +0200 | |
changeset 82697 | cc05bc2cfb2f |
parent 82663 | bd951e02d6b9 |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/HOL.thy |
11750 | 2 |
Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson |
3 |
*) |
|
923 | 4 |
|
60758 | 5 |
section \<open>The basis of Higher-Order Logic\<close> |
923 | 6 |
|
15131 | 7 |
theory HOL |
82360 | 8 |
imports Pure Try0 Tools.Code_Generator |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46497
diff
changeset
|
9 |
keywords |
52432 | 10 |
"try" "solve_direct" "quickcheck" "print_coercions" "print_claset" |
11 |
"print_induct_rules" :: diag and |
|
47657
1ba213363d0c
moved modules with only vague relation to the code generator to theory HOL rather than theory Code_Generator
haftmann
parents:
46973
diff
changeset
|
12 |
"quickcheck_params" :: thy_decl |
71833
ff6f3b09b8b4
abbrevs for the Uniq quantifier; trying Sup_nat_def to allow zero (experimentally)
paulson <lp15@cam.ac.uk>
parents:
71827
diff
changeset
|
13 |
abbrevs "?<" = "\<exists>\<^sub>\<le>\<^sub>1" |
15131 | 14 |
begin |
2260 | 15 |
|
69605 | 16 |
ML_file \<open>~~/src/Tools/misc_legacy.ML\<close> |
17 |
ML_file \<open>~~/src/Tools/try.ML\<close> |
|
18 |
ML_file \<open>~~/src/Tools/quickcheck.ML\<close> |
|
19 |
ML_file \<open>~~/src/Tools/solve_direct.ML\<close> |
|
20 |
ML_file \<open>~~/src/Tools/IsaPlanner/zipper.ML\<close> |
|
21 |
ML_file \<open>~~/src/Tools/IsaPlanner/isand.ML\<close> |
|
22 |
ML_file \<open>~~/src/Tools/IsaPlanner/rw_inst.ML\<close> |
|
23 |
ML_file \<open>~~/src/Provers/hypsubst.ML\<close> |
|
24 |
ML_file \<open>~~/src/Provers/splitter.ML\<close> |
|
25 |
ML_file \<open>~~/src/Provers/classical.ML\<close> |
|
26 |
ML_file \<open>~~/src/Provers/blast.ML\<close> |
|
27 |
ML_file \<open>~~/src/Provers/clasimp.ML\<close> |
|
28 |
ML_file \<open>~~/src/Tools/eqsubst.ML\<close> |
|
29 |
ML_file \<open>~~/src/Provers/quantifier1.ML\<close> |
|
30 |
ML_file \<open>~~/src/Tools/atomize_elim.ML\<close> |
|
31 |
ML_file \<open>~~/src/Tools/cong_tac.ML\<close> |
|
32 |
ML_file \<open>~~/src/Tools/intuitionistic.ML\<close> setup \<open>Intuitionistic.method_setup \<^binding>\<open>iprover\<close>\<close> |
|
33 |
ML_file \<open>~~/src/Tools/project_rule.ML\<close> |
|
34 |
ML_file \<open>~~/src/Tools/subtyping.ML\<close> |
|
35 |
ML_file \<open>~~/src/Tools/case_product.ML\<close> |
|
48891 | 36 |
|
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82360
diff
changeset
|
37 |
ML \<open> |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82360
diff
changeset
|
38 |
val _ = |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82360
diff
changeset
|
39 |
Try.tool_setup |
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82360
diff
changeset
|
40 |
{name = "try0", weight = 30, auto_option = \<^system_option>\<open>auto_methods\<close>, |
82364 | 41 |
body = fn auto => fst o Try0.generic_try0 (if auto then Try0.Auto_Try else Try0.Try) NONE Try0.empty_facts} |
82363
3a7fc54b50ca
tuned and moved configuration of auto_try0 to theory HOL
desharna
parents:
82360
diff
changeset
|
42 |
\<close> |
30165
6ee87f67d9cd
moved generic intuitionistic prover to src/Tools/intuitionistic.ML;
wenzelm
parents:
30160
diff
changeset
|
43 |
|
67149 | 44 |
ML \<open>Plugin_Name.declare_setup \<^binding>\<open>extraction\<close>\<close> |
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
45 |
|
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
46 |
ML \<open> |
67149 | 47 |
Plugin_Name.declare_setup \<^binding>\<open>quickcheck_random\<close>; |
48 |
Plugin_Name.declare_setup \<^binding>\<open>quickcheck_exhaustive\<close>; |
|
49 |
Plugin_Name.declare_setup \<^binding>\<open>quickcheck_bounded_forall\<close>; |
|
50 |
Plugin_Name.declare_setup \<^binding>\<open>quickcheck_full_exhaustive\<close>; |
|
51 |
Plugin_Name.declare_setup \<^binding>\<open>quickcheck_narrowing\<close>; |
|
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
52 |
\<close> |
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
53 |
ML \<open> |
67149 | 54 |
Plugin_Name.define_setup \<^binding>\<open>quickcheck\<close> |
69593 | 55 |
[\<^plugin>\<open>quickcheck_exhaustive\<close>, |
56 |
\<^plugin>\<open>quickcheck_random\<close>, |
|
57 |
\<^plugin>\<open>quickcheck_bounded_forall\<close>, |
|
58 |
\<^plugin>\<open>quickcheck_full_exhaustive\<close>, |
|
59 |
\<^plugin>\<open>quickcheck_narrowing\<close>] |
|
58659
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
60 |
\<close> |
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
61 |
|
6c9821c32dd5
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
wenzelm
parents:
57964
diff
changeset
|
62 |
|
60758 | 63 |
subsection \<open>Primitive logic\<close> |
11750 | 64 |
|
66893
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
65 |
text \<open> |
77172 | 66 |
The definition of the logic is based on Mike Gordon's technical report \<^cite>\<open>"Gordon-TR68"\<close> that |
66893
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
67 |
describes the first implementation of HOL. However, there are a number of differences. |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
68 |
In particular, we start with the definite description operator and introduce Hilbert's \<open>\<epsilon>\<close> operator |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
69 |
only much later. Moreover, axiom \<open>(P \<longrightarrow> Q) \<longrightarrow> (Q \<longrightarrow> P) \<longrightarrow> (P = Q)\<close> is derived from the other |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
70 |
axioms. The fact that this axiom is derivable was first noticed by Bruno Barras (for Mike Gordon's |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
71 |
line of HOL systems) and later independently by Alexander Maletzky (for Isabelle/HOL). |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
72 |
\<close> |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
73 |
|
60758 | 74 |
subsubsection \<open>Core syntax\<close> |
2260 | 75 |
|
67149 | 76 |
setup \<open>Axclass.class_axiomatization (\<^binding>\<open>type\<close>, [])\<close> |
36452 | 77 |
default_sort type |
69593 | 78 |
setup \<open>Object_Logic.add_base_sort \<^sort>\<open>type\<close>\<close> |
25460
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
79 |
|
70879 | 80 |
setup \<open>Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\<close> |
70849 | 81 |
|
55383 | 82 |
axiomatization where fun_arity: "OFCLASS('a \<Rightarrow> 'b, type_class)" |
83 |
instance "fun" :: (type, type) type by (rule fun_arity) |
|
84 |
||
85 |
axiomatization where itself_arity: "OFCLASS('a itself, type_class)" |
|
86 |
instance itself :: (type) type by (rule itself_arity) |
|
25460
b80087af2274
interpretation of typedecls: instantiation to class type
haftmann
parents:
25388
diff
changeset
|
87 |
|
7357 | 88 |
typedecl bool |
923 | 89 |
|
80934 | 90 |
judgment Trueprop :: "bool \<Rightarrow> prop" (\<open>(\<open>notation=judgment\<close>_)\<close> 5) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
91 |
|
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
92 |
axiomatization implies :: "[bool, bool] \<Rightarrow> bool" (infixr \<open>\<longrightarrow>\<close> 25) |
68973
a1e26440efb8
tuned "=" syntax declarations; made "~=" uniformly "infix"
nipkow
parents:
68072
diff
changeset
|
93 |
and eq :: "['a, 'a] \<Rightarrow> bool" |
62151 | 94 |
and The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a" |
95 |
||
68973
a1e26440efb8
tuned "=" syntax declarations; made "~=" uniformly "infix"
nipkow
parents:
68072
diff
changeset
|
96 |
notation (input) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
97 |
eq (infixl \<open>=\<close> 50) |
68973
a1e26440efb8
tuned "=" syntax declarations; made "~=" uniformly "infix"
nipkow
parents:
68072
diff
changeset
|
98 |
notation (output) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
99 |
eq (infix \<open>=\<close> 50) |
68973
a1e26440efb8
tuned "=" syntax declarations; made "~=" uniformly "infix"
nipkow
parents:
68072
diff
changeset
|
100 |
|
a1e26440efb8
tuned "=" syntax declarations; made "~=" uniformly "infix"
nipkow
parents:
68072
diff
changeset
|
101 |
text \<open>The input syntax for \<open>eq\<close> is more permissive than the output syntax |
a1e26440efb8
tuned "=" syntax declarations; made "~=" uniformly "infix"
nipkow
parents:
68072
diff
changeset
|
102 |
because of the large amount of material that relies on infixl.\<close> |
923 | 103 |
|
62151 | 104 |
subsubsection \<open>Defined connectives and quantifiers\<close> |
105 |
||
106 |
definition True :: bool |
|
107 |
where "True \<equiv> ((\<lambda>x::bool. x) = (\<lambda>x. x))" |
|
108 |
||
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
109 |
definition All :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder \<open>\<forall>\<close> 10) |
62151 | 110 |
where "All P \<equiv> (P = (\<lambda>x. True))" |
46973 | 111 |
|
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
112 |
definition Ex :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder \<open>\<exists>\<close> 10) |
62151 | 113 |
where "Ex P \<equiv> \<forall>Q. (\<forall>x. P x \<longrightarrow> Q) \<longrightarrow> Q" |
114 |
||
115 |
definition False :: bool |
|
116 |
where "False \<equiv> (\<forall>P. P)" |
|
117 |
||
81125 | 118 |
definition Not :: "bool \<Rightarrow> bool" (\<open>(\<open>open_block notation=\<open>prefix \<not>\<close>\<close>\<not> _)\<close> [40] 40) |
62151 | 119 |
where not_def: "\<not> P \<equiv> P \<longrightarrow> False" |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
120 |
|
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
121 |
definition conj :: "[bool, bool] \<Rightarrow> bool" (infixr \<open>\<and>\<close> 35) |
62151 | 122 |
where and_def: "P \<and> Q \<equiv> \<forall>R. (P \<longrightarrow> Q \<longrightarrow> R) \<longrightarrow> R" |
38555 | 123 |
|
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
124 |
definition disj :: "[bool, bool] \<Rightarrow> bool" (infixr \<open>\<or>\<close> 30) |
62151 | 125 |
where or_def: "P \<or> Q \<equiv> \<forall>R. (P \<longrightarrow> R) \<longrightarrow> (Q \<longrightarrow> R) \<longrightarrow> R" |
126 |
||
71827 | 127 |
definition Uniq :: "('a \<Rightarrow> bool) \<Rightarrow> bool" |
128 |
where "Uniq P \<equiv> (\<forall>x y. P x \<longrightarrow> P y \<longrightarrow> y = x)" |
|
129 |
||
63909 | 130 |
definition Ex1 :: "('a \<Rightarrow> bool) \<Rightarrow> bool" |
62151 | 131 |
where "Ex1 P \<equiv> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> y = x)" |
923 | 132 |
|
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
133 |
|
60758 | 134 |
subsubsection \<open>Additional concrete syntax\<close> |
2260 | 135 |
|
80934 | 136 |
syntax (ASCII) "_Uniq" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=4 notation=\<open>binder ?<\<close>\<close>?< _./ _)\<close> [0, 10] 10) |
137 |
syntax "_Uniq" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=2 notation=\<open>binder \<exists>\<^sub>\<le>\<^sub>1\<close>\<close>\<exists>\<^sub>\<le>\<^sub>1 _./ _)\<close> [0, 10] 10) |
|
80760 | 138 |
|
139 |
syntax_consts "_Uniq" \<rightleftharpoons> Uniq |
|
140 |
||
71827 | 141 |
translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)" |
142 |
||
81202 | 143 |
typed_print_translation \<open> |
81545
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81202
diff
changeset
|
144 |
[(\<^const_syntax>\<open>Uniq\<close>, Syntax_Trans.preserve_binder_abs_tr' \<^syntax_const>\<open>_Uniq\<close>)] |
71827 | 145 |
\<close> \<comment> \<open>to avoid eta-contraction of body\<close> |
146 |
||
147 |
||
63909 | 148 |
syntax (ASCII) |
80934 | 149 |
"_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=3 notation=\<open>binder EX!\<close>\<close>EX! _./ _)\<close> [0, 10] 10) |
63909 | 150 |
syntax (input) |
80934 | 151 |
"_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=3 notation=\<open>binder ?!\<close>\<close>?! _./ _)\<close> [0, 10] 10) |
152 |
syntax "_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=3 notation=\<open>binder \<exists>!\<close>\<close>\<exists>!_./ _)\<close> [0, 10] 10) |
|
80760 | 153 |
|
154 |
syntax_consts "_Ex1" \<rightleftharpoons> Ex1 |
|
155 |
||
63909 | 156 |
translations "\<exists>!x. P" \<rightleftharpoons> "CONST Ex1 (\<lambda>x. P)" |
157 |
||
81202 | 158 |
typed_print_translation \<open> |
81545
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81202
diff
changeset
|
159 |
[(\<^const_syntax>\<open>Ex1\<close>, Syntax_Trans.preserve_binder_abs_tr' \<^syntax_const>\<open>_Ex1\<close>)] |
63909 | 160 |
\<close> \<comment> \<open>to avoid eta-contraction of body\<close> |
161 |
||
162 |
||
63912
9f8325206465
clarified notation: iterated quantifier is negated as one chunk;
wenzelm
parents:
63909
diff
changeset
|
163 |
syntax |
80934 | 164 |
"_Not_Ex" :: "idts \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=3 notation=\<open>binder \<nexists>\<close>\<close>\<nexists>_./ _)\<close> [0, 10] 10) |
165 |
"_Not_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" (\<open>(\<open>indent=3 notation=\<open>binder \<nexists>!\<close>\<close>\<nexists>!_./ _)\<close> [0, 10] 10) |
|
80760 | 166 |
syntax_consts |
167 |
"_Not_Ex" \<rightleftharpoons> Ex and |
|
168 |
"_Not_Ex1" \<rightleftharpoons> Ex1 |
|
63912
9f8325206465
clarified notation: iterated quantifier is negated as one chunk;
wenzelm
parents:
63909
diff
changeset
|
169 |
translations |
9f8325206465
clarified notation: iterated quantifier is negated as one chunk;
wenzelm
parents:
63909
diff
changeset
|
170 |
"\<nexists>x. P" \<rightleftharpoons> "\<not> (\<exists>x. P)" |
9f8325206465
clarified notation: iterated quantifier is negated as one chunk;
wenzelm
parents:
63909
diff
changeset
|
171 |
"\<nexists>!x. P" \<rightleftharpoons> "\<not> (\<exists>!x. P)" |
62522 | 172 |
|
173 |
||
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
174 |
abbreviation not_equal :: "['a, 'a] \<Rightarrow> bool" (infix \<open>\<noteq>\<close> 50) |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
175 |
where "x \<noteq> y \<equiv> \<not> (x = y)" |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
176 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
177 |
notation (ASCII) |
81125 | 178 |
Not (\<open>(\<open>open_block notation=\<open>prefix ~\<close>\<close>~ _)\<close> [40] 40) and |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
179 |
conj (infixr \<open>&\<close> 35) and |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
180 |
disj (infixr \<open>|\<close> 30) and |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
181 |
implies (infixr \<open>-->\<close> 25) and |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
182 |
not_equal (infix \<open>~=\<close> 50) |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
183 |
|
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
184 |
abbreviation (iff) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
185 |
iff :: "[bool, bool] \<Rightarrow> bool" (infixr \<open>\<longleftrightarrow>\<close> 25) |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
186 |
where "A \<longleftrightarrow> B \<equiv> A = B" |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19607
diff
changeset
|
187 |
|
80934 | 188 |
syntax "_The" :: "[pttrn, bool] \<Rightarrow> 'a" (\<open>(\<open>indent=3 notation=\<open>binder THE\<close>\<close>THE _./ _)\<close> [0, 10] 10) |
80760 | 189 |
syntax_consts "_The" \<rightleftharpoons> The |
60759 | 190 |
translations "THE x. P" \<rightleftharpoons> "CONST The (\<lambda>x. P)" |
60758 | 191 |
print_translation \<open> |
81545
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81202
diff
changeset
|
192 |
[(\<^const_syntax>\<open>The\<close>, fn ctxt => fn [Abs abs] => |
6f8a56a6b391
clarified renaming of bounds, using Syntax_Trans.variant_bounds: avoid structures and fixed variables with syntax;
wenzelm
parents:
81202
diff
changeset
|
193 |
let val (x, t) = Syntax_Trans.atomic_abs_tr' ctxt abs |
69593 | 194 |
in Syntax.const \<^syntax_const>\<open>_The\<close> $ x $ t end)] |
61799 | 195 |
\<close> \<comment> \<open>To avoid eta-contraction of body\<close> |
923 | 196 |
|
46125
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
45654
diff
changeset
|
197 |
nonterminal case_syn and cases_syn |
00cd193a48dc
improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents:
45654
diff
changeset
|
198 |
syntax |
80934 | 199 |
"_case_syntax" :: "['a, cases_syn] \<Rightarrow> 'b" (\<open>(\<open>notation=\<open>mixfix case expression\<close>\<close>case _ of/ _)\<close> 10) |
81125 | 200 |
"_case1" :: "['a, 'b] \<Rightarrow> case_syn" |
201 |
(\<open>(\<open>indent=2 notation=\<open>mixfix case clause\<close>\<close>(\<open>open_block notation=\<open>pattern case\<close>\<close>_) \<Rightarrow>/ _)\<close> 10) |
|
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
202 |
"" :: "case_syn \<Rightarrow> cases_syn" (\<open>_\<close>) |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
203 |
"_case2" :: "[case_syn, cases_syn] \<Rightarrow> cases_syn" (\<open>_/ | _\<close>) |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
204 |
syntax (ASCII) |
81125 | 205 |
"_case1" :: "['a, 'b] \<Rightarrow> case_syn" |
206 |
(\<open>(\<open>indent=2 notation=\<open>mixfix case clause\<close>\<close>(\<open>open_block notation=\<open>pattern case\<close>\<close>_) =>/ _)\<close> 10) |
|
13763
f94b569cd610
added print translations tha avoid eta contraction for important binders.
nipkow
parents:
13723
diff
changeset
|
207 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61941
diff
changeset
|
208 |
notation (ASCII) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
209 |
All (binder \<open>ALL \<close> 10) and |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
210 |
Ex (binder \<open>EX \<close> 10) |
2372 | 211 |
|
62521 | 212 |
notation (input) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
213 |
All (binder \<open>! \<close> 10) and |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
214 |
Ex (binder \<open>? \<close> 10) |
7238
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
215 |
|
36e58620ffc8
replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents:
7220
diff
changeset
|
216 |
|
60758 | 217 |
subsubsection \<open>Axioms and basic definitions\<close> |
2260 | 218 |
|
46973 | 219 |
axiomatization where |
220 |
refl: "t = (t::'a)" and |
|
221 |
subst: "s = t \<Longrightarrow> P s \<Longrightarrow> P t" and |
|
60759 | 222 |
ext: "(\<And>x::'a. (f x ::'b) = g x) \<Longrightarrow> (\<lambda>x. f x) = (\<lambda>x. g x)" |
61799 | 223 |
\<comment> \<open>Extensionality is built into the meta-logic, and this rule expresses |
15380 | 224 |
a related property. It is an eta-expanded version of the traditional |
60758 | 225 |
rule, and similar to the ABS rule of HOL\<close> and |
6289 | 226 |
|
11432
8a203ae6efe3
added "The" (definite description operator) (by Larry);
wenzelm
parents:
10489
diff
changeset
|
227 |
the_eq_trivial: "(THE x. x = a) = (a::'a)" |
923 | 228 |
|
46973 | 229 |
axiomatization where |
60759 | 230 |
impI: "(P \<Longrightarrow> Q) \<Longrightarrow> P \<longrightarrow> Q" and |
231 |
mp: "\<lbrakk>P \<longrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q" and |
|
15380 | 232 |
|
60759 | 233 |
True_or_False: "(P = True) \<or> (P = False)" |
15380 | 234 |
|
80934 | 235 |
definition If :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(\<open>notation=\<open>mixfix if expression\<close>\<close>if (_)/ then (_)/ else (_))\<close> [0, 0, 10] 10) |
60759 | 236 |
where "If P x y \<equiv> (THE z::'a. (P = True \<longrightarrow> z = x) \<and> (P = False \<longrightarrow> z = y))" |
923 | 237 |
|
46973 | 238 |
definition Let :: "'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b" |
239 |
where "Let s f \<equiv> f s" |
|
38525 | 240 |
|
80762 | 241 |
nonterminal letbinds and letbind |
81595
ed264056f5dc
more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents:
81545
diff
changeset
|
242 |
|
ed264056f5dc
more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents:
81545
diff
changeset
|
243 |
open_bundle let_syntax |
ed264056f5dc
more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents:
81545
diff
changeset
|
244 |
begin |
ed264056f5dc
more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents:
81545
diff
changeset
|
245 |
|
80762 | 246 |
syntax |
80934 | 247 |
"_bind" :: "[pttrn, 'a] \<Rightarrow> letbind" (\<open>(\<open>indent=2 notation=\<open>mixfix let binding\<close>\<close>_ =/ _)\<close> 10) |
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
248 |
"" :: "letbind \<Rightarrow> letbinds" (\<open>_\<close>) |
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
249 |
"_binds" :: "[letbind, letbinds] \<Rightarrow> letbinds" (\<open>_;/ _\<close>) |
80934 | 250 |
"_Let" :: "[letbinds, 'a] \<Rightarrow> 'a" (\<open>(\<open>notation=\<open>mixfix let expression\<close>\<close>let (_)/ in (_))\<close> [0, 10] 10) |
80760 | 251 |
syntax_consts |
252 |
"_bind" "_binds" "_Let" \<rightleftharpoons> Let |
|
38525 | 253 |
translations |
60759 | 254 |
"_Let (_binds b bs) e" \<rightleftharpoons> "_Let b (_Let bs e)" |
255 |
"let x = a in e" \<rightleftharpoons> "CONST Let a (\<lambda>x. e)" |
|
38525 | 256 |
|
81595
ed264056f5dc
more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents:
81545
diff
changeset
|
257 |
end |
ed264056f5dc
more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents:
81545
diff
changeset
|
258 |
|
46973 | 259 |
axiomatization undefined :: 'a |
22481
79c2724c36b5
added class "default" and expansion axioms for undefined
haftmann
parents:
22473
diff
changeset
|
260 |
|
46973 | 261 |
class default = fixes default :: 'a |
4868 | 262 |
|
11750 | 263 |
|
60758 | 264 |
subsection \<open>Fundamental rules\<close> |
20944 | 265 |
|
60758 | 266 |
subsubsection \<open>Equality\<close> |
20944 | 267 |
|
60759 | 268 |
lemma sym: "s = t \<Longrightarrow> t = s" |
18457 | 269 |
by (erule subst) (rule refl) |
15411 | 270 |
|
60759 | 271 |
lemma ssubst: "t = s \<Longrightarrow> P s \<Longrightarrow> P t" |
18457 | 272 |
by (drule sym) (erule subst) |
15411 | 273 |
|
60759 | 274 |
lemma trans: "\<lbrakk>r = s; s = t\<rbrakk> \<Longrightarrow> r = t" |
18457 | 275 |
by (erule subst) |
15411 | 276 |
|
60759 | 277 |
lemma trans_sym [Pure.elim?]: "r = s \<Longrightarrow> t = s \<Longrightarrow> r = t" |
40715
3ba17f07b23c
lemma trans_sym allows single-step "normalization" in Isar, e.g. via moreover/ultimately;
wenzelm
parents:
40582
diff
changeset
|
278 |
by (rule trans [OF _ sym]) |
3ba17f07b23c
lemma trans_sym allows single-step "normalization" in Isar, e.g. via moreover/ultimately;
wenzelm
parents:
40582
diff
changeset
|
279 |
|
58826 | 280 |
lemma meta_eq_to_obj_eq: |
63575 | 281 |
assumes "A \<equiv> B" |
20944 | 282 |
shows "A = B" |
63575 | 283 |
unfolding assms by (rule refl) |
15411 | 284 |
|
61799 | 285 |
text \<open>Useful with \<open>erule\<close> for proving equalities from known equalities.\<close> |
20944 | 286 |
(* a = b |
15411 | 287 |
| | |
288 |
c = d *) |
|
60759 | 289 |
lemma box_equals: "\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
290 |
by (iprover intro: sym trans) |
15411 | 291 |
|
60758 | 292 |
text \<open>For calculational reasoning:\<close> |
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
293 |
|
60759 | 294 |
lemma forw_subst: "a = b \<Longrightarrow> P b \<Longrightarrow> P a" |
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
295 |
by (rule ssubst) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
296 |
|
60759 | 297 |
lemma back_subst: "P a \<Longrightarrow> a = b \<Longrightarrow> P b" |
15524
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
298 |
by (rule subst) |
2ef571f80a55
Moved oderings from HOL into the new Orderings.thy
nipkow
parents:
15481
diff
changeset
|
299 |
|
15411 | 300 |
|
60758 | 301 |
subsubsection \<open>Congruence rules for application\<close> |
15411 | 302 |
|
61799 | 303 |
text \<open>Similar to \<open>AP_THM\<close> in Gordon's HOL.\<close> |
60759 | 304 |
lemma fun_cong: "(f :: 'a \<Rightarrow> 'b) = g \<Longrightarrow> f x = g x" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
305 |
by (iprover intro: refl elim: subst) |
15411 | 306 |
|
61799 | 307 |
text \<open>Similar to \<open>AP_TERM\<close> in Gordon's HOL and FOL's \<open>subst_context\<close>.\<close> |
60759 | 308 |
lemma arg_cong: "x = y \<Longrightarrow> f x = f y" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
309 |
by (iprover intro: refl elim: subst) |
15411 | 310 |
|
60759 | 311 |
lemma arg_cong2: "\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> f a c = f b d" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
312 |
by (iprover intro: refl elim: subst) |
15655 | 313 |
|
60759 | 314 |
lemma cong: "\<lbrakk>f = g; (x::'a) = y\<rbrakk> \<Longrightarrow> f x = g y" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
315 |
by (iprover intro: refl elim: subst) |
15411 | 316 |
|
60758 | 317 |
ML \<open>fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\<close> |
15411 | 318 |
|
32733
71618deaf777
moved generic cong_tac from HOL/Tools/datatype_aux.ML to Tools/cong_tac.ML, proper subgoal selection (failure, not exception);
wenzelm
parents:
32668
diff
changeset
|
319 |
|
60758 | 320 |
subsubsection \<open>Equality of booleans -- iff\<close> |
15411 | 321 |
|
60759 | 322 |
lemma iffD2: "\<lbrakk>P = Q; Q\<rbrakk> \<Longrightarrow> P" |
18457 | 323 |
by (erule ssubst) |
15411 | 324 |
|
60759 | 325 |
lemma rev_iffD2: "\<lbrakk>Q; P = Q\<rbrakk> \<Longrightarrow> P" |
18457 | 326 |
by (erule iffD2) |
15411 | 327 |
|
21504 | 328 |
lemma iffD1: "Q = P \<Longrightarrow> Q \<Longrightarrow> P" |
329 |
by (drule sym) (rule iffD2) |
|
330 |
||
331 |
lemma rev_iffD1: "Q \<Longrightarrow> Q = P \<Longrightarrow> P" |
|
332 |
by (drule sym) (rule rev_iffD2) |
|
15411 | 333 |
|
334 |
lemma iffE: |
|
60759 | 335 |
assumes major: "P = Q" |
336 |
and minor: "\<lbrakk>P \<longrightarrow> Q; Q \<longrightarrow> P\<rbrakk> \<Longrightarrow> R" |
|
18457 | 337 |
shows R |
338 |
by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1]) |
|
15411 | 339 |
|
340 |
||
66893
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
341 |
subsubsection \<open>True (1)\<close> |
15411 | 342 |
|
63575 | 343 |
lemma TrueI: True |
21504 | 344 |
unfolding True_def by (rule refl) |
15411 | 345 |
|
60759 | 346 |
lemma eqTrueE: "P = True \<Longrightarrow> P" |
21504 | 347 |
by (erule iffD2) (rule TrueI) |
15411 | 348 |
|
349 |
||
66893
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
350 |
subsubsection \<open>Universal quantifier (1)\<close> |
15411 | 351 |
|
60759 | 352 |
lemma spec: "\<forall>x::'a. P x \<Longrightarrow> P x" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
353 |
unfolding All_def by (iprover intro: eqTrueE fun_cong) |
15411 | 354 |
|
355 |
lemma allE: |
|
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
356 |
assumes major: "\<forall>x. P x" and minor: "P x \<Longrightarrow> R" |
21504 | 357 |
shows R |
358 |
by (iprover intro: minor major [THEN spec]) |
|
15411 | 359 |
|
360 |
lemma all_dupE: |
|
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
361 |
assumes major: "\<forall>x. P x" and minor: "\<lbrakk>P x; \<forall>x. P x\<rbrakk> \<Longrightarrow> R" |
21504 | 362 |
shows R |
363 |
by (iprover intro: minor major major [THEN spec]) |
|
15411 | 364 |
|
365 |
||
60758 | 366 |
subsubsection \<open>False\<close> |
21504 | 367 |
|
60758 | 368 |
text \<open> |
61799 | 369 |
Depends upon \<open>spec\<close>; it is impossible to do propositional |
21504 | 370 |
logic before quantifiers! |
60758 | 371 |
\<close> |
15411 | 372 |
|
60759 | 373 |
lemma FalseE: "False \<Longrightarrow> P" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
374 |
unfolding False_def by (erule spec) |
15411 | 375 |
|
60759 | 376 |
lemma False_neq_True: "False = True \<Longrightarrow> P" |
21504 | 377 |
by (erule eqTrueE [THEN FalseE]) |
15411 | 378 |
|
379 |
||
60758 | 380 |
subsubsection \<open>Negation\<close> |
15411 | 381 |
|
382 |
lemma notI: |
|
60759 | 383 |
assumes "P \<Longrightarrow> False" |
384 |
shows "\<not> P" |
|
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
385 |
unfolding not_def by (iprover intro: impI assms) |
15411 | 386 |
|
60759 | 387 |
lemma False_not_True: "False \<noteq> True" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
388 |
by (iprover intro: notI elim: False_neq_True) |
15411 | 389 |
|
60759 | 390 |
lemma True_not_False: "True \<noteq> False" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
391 |
by (iprover intro: notI dest: sym elim: False_neq_True) |
15411 | 392 |
|
60759 | 393 |
lemma notE: "\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
394 |
unfolding not_def |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
395 |
by (iprover intro: mp [THEN FalseE]) |
15411 | 396 |
|
397 |
||
60758 | 398 |
subsubsection \<open>Implication\<close> |
15411 | 399 |
|
400 |
lemma impE: |
|
60759 | 401 |
assumes "P \<longrightarrow> Q" P "Q \<Longrightarrow> R" |
402 |
shows R |
|
63575 | 403 |
by (iprover intro: assms mp) |
15411 | 404 |
|
63575 | 405 |
text \<open>Reduces \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, allowing substitution in \<open>P\<close>.\<close> |
60759 | 406 |
lemma rev_mp: "\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
407 |
by (rule mp) |
15411 | 408 |
|
409 |
lemma contrapos_nn: |
|
60759 | 410 |
assumes major: "\<not> Q" |
63575 | 411 |
and minor: "P \<Longrightarrow> Q" |
60759 | 412 |
shows "\<not> P" |
63575 | 413 |
by (iprover intro: notI minor major [THEN notE]) |
15411 | 414 |
|
63575 | 415 |
text \<open>Not used at all, but we already have the other 3 combinations.\<close> |
15411 | 416 |
lemma contrapos_pn: |
417 |
assumes major: "Q" |
|
63575 | 418 |
and minor: "P \<Longrightarrow> \<not> Q" |
60759 | 419 |
shows "\<not> P" |
63575 | 420 |
by (iprover intro: notI minor major notE) |
15411 | 421 |
|
60759 | 422 |
lemma not_sym: "t \<noteq> s \<Longrightarrow> s \<noteq> t" |
21250 | 423 |
by (erule contrapos_nn) (erule sym) |
424 |
||
60759 | 425 |
lemma eq_neq_eq_imp_neq: "\<lbrakk>x = a; a \<noteq> b; b = y\<rbrakk> \<Longrightarrow> x \<noteq> y" |
21250 | 426 |
by (erule subst, erule ssubst, assumption) |
15411 | 427 |
|
428 |
||
66893
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
429 |
subsubsection \<open>Disjunction (1)\<close> |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
430 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
431 |
lemma disjE: |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
432 |
assumes major: "P \<or> Q" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
433 |
and minorP: "P \<Longrightarrow> R" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
434 |
and minorQ: "Q \<Longrightarrow> R" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
435 |
shows R |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
436 |
by (iprover intro: minorP minorQ impI |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
437 |
major [unfolded or_def, THEN spec, THEN mp, THEN mp]) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
438 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
439 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
440 |
subsubsection \<open>Derivation of \<open>iffI\<close>\<close> |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
441 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
442 |
text \<open>In an intuitionistic version of HOL \<open>iffI\<close> needs to be an axiom.\<close> |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
443 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
444 |
lemma iffI: |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
445 |
assumes "P \<Longrightarrow> Q" and "Q \<Longrightarrow> P" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
446 |
shows "P = Q" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
447 |
proof (rule disjE[OF True_or_False[of P]]) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
448 |
assume 1: "P = True" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
449 |
note Q = assms(1)[OF eqTrueE[OF this]] |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
450 |
from 1 show ?thesis |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
451 |
proof (rule ssubst) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
452 |
from True_or_False[of Q] show "True = Q" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
453 |
proof (rule disjE) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
454 |
assume "Q = True" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
455 |
thus ?thesis by(rule sym) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
456 |
next |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
457 |
assume "Q = False" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
458 |
with Q have False by (rule rev_iffD1) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
459 |
thus ?thesis by (rule FalseE) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
460 |
qed |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
461 |
qed |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
462 |
next |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
463 |
assume 2: "P = False" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
464 |
thus ?thesis |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
465 |
proof (rule ssubst) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
466 |
from True_or_False[of Q] show "False = Q" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
467 |
proof (rule disjE) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
468 |
assume "Q = True" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
469 |
from 2 assms(2)[OF eqTrueE[OF this]] have False by (rule iffD1) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
470 |
thus ?thesis by (rule FalseE) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
471 |
next |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
472 |
assume "Q = False" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
473 |
thus ?thesis by(rule sym) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
474 |
qed |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
475 |
qed |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
476 |
qed |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
477 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
478 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
479 |
subsubsection \<open>True (2)\<close> |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
480 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
481 |
lemma eqTrueI: "P \<Longrightarrow> P = True" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
482 |
by (iprover intro: iffI TrueI) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
483 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
484 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
485 |
subsubsection \<open>Universal quantifier (2)\<close> |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
486 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
487 |
lemma allI: |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
488 |
assumes "\<And>x::'a. P x" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
489 |
shows "\<forall>x. P x" |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
490 |
unfolding All_def by (iprover intro: ext eqTrueI assms) |
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
491 |
|
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
492 |
|
60758 | 493 |
subsubsection \<open>Existential quantifier\<close> |
15411 | 494 |
|
60759 | 495 |
lemma exI: "P x \<Longrightarrow> \<exists>x::'a. P x" |
63575 | 496 |
unfolding Ex_def by (iprover intro: allI allE impI mp) |
15411 | 497 |
|
498 |
lemma exE: |
|
60759 | 499 |
assumes major: "\<exists>x::'a. P x" |
63575 | 500 |
and minor: "\<And>x. P x \<Longrightarrow> Q" |
15411 | 501 |
shows "Q" |
63575 | 502 |
by (rule major [unfolded Ex_def, THEN spec, THEN mp]) (iprover intro: impI [THEN allI] minor) |
15411 | 503 |
|
504 |
||
60758 | 505 |
subsubsection \<open>Conjunction\<close> |
15411 | 506 |
|
60759 | 507 |
lemma conjI: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q" |
63575 | 508 |
unfolding and_def by (iprover intro: impI [THEN allI] mp) |
15411 | 509 |
|
60759 | 510 |
lemma conjunct1: "\<lbrakk>P \<and> Q\<rbrakk> \<Longrightarrow> P" |
63575 | 511 |
unfolding and_def by (iprover intro: impI dest: spec mp) |
15411 | 512 |
|
60759 | 513 |
lemma conjunct2: "\<lbrakk>P \<and> Q\<rbrakk> \<Longrightarrow> Q" |
63575 | 514 |
unfolding and_def by (iprover intro: impI dest: spec mp) |
15411 | 515 |
|
516 |
lemma conjE: |
|
60759 | 517 |
assumes major: "P \<and> Q" |
63575 | 518 |
and minor: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R" |
60759 | 519 |
shows R |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
520 |
proof (rule minor) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
521 |
show P by (rule major [THEN conjunct1]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
522 |
show Q by (rule major [THEN conjunct2]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
523 |
qed |
15411 | 524 |
|
525 |
lemma context_conjI: |
|
63575 | 526 |
assumes P "P \<Longrightarrow> Q" |
527 |
shows "P \<and> Q" |
|
528 |
by (iprover intro: conjI assms) |
|
15411 | 529 |
|
530 |
||
66893
ced164fe3bbd
derived axiom iffI as a lemma (thanks to Alexander Maletzky)
nipkow
parents:
66836
diff
changeset
|
531 |
subsubsection \<open>Disjunction (2)\<close> |
15411 | 532 |
|
60759 | 533 |
lemma disjI1: "P \<Longrightarrow> P \<or> Q" |
63575 | 534 |
unfolding or_def by (iprover intro: allI impI mp) |
15411 | 535 |
|
60759 | 536 |
lemma disjI2: "Q \<Longrightarrow> P \<or> Q" |
63575 | 537 |
unfolding or_def by (iprover intro: allI impI mp) |
15411 | 538 |
|
539 |
||
60758 | 540 |
subsubsection \<open>Classical logic\<close> |
15411 | 541 |
|
542 |
lemma classical: |
|
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
543 |
assumes "\<not> P \<Longrightarrow> P" |
60759 | 544 |
shows P |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
545 |
proof (rule True_or_False [THEN disjE]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
546 |
show P if "P = True" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
547 |
using that by (iprover intro: eqTrueE) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
548 |
show P if "P = False" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
549 |
proof (intro notI assms) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
550 |
assume P |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
551 |
with that show False |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
552 |
by (iprover elim: subst) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
553 |
qed |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
554 |
qed |
15411 | 555 |
|
45607 | 556 |
lemmas ccontr = FalseE [THEN classical] |
15411 | 557 |
|
63575 | 558 |
text \<open>\<open>notE\<close> with premises exchanged; it discharges \<open>\<not> R\<close> so that it can be used to |
559 |
make elimination rules.\<close> |
|
15411 | 560 |
lemma rev_notE: |
60759 | 561 |
assumes premp: P |
63575 | 562 |
and premnot: "\<not> R \<Longrightarrow> \<not> P" |
60759 | 563 |
shows R |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
564 |
by (iprover intro: ccontr notE [OF premnot premp]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
565 |
|
15411 | 566 |
|
63575 | 567 |
text \<open>Double negation law.\<close> |
60759 | 568 |
lemma notnotD: "\<not>\<not> P \<Longrightarrow> P" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
569 |
by (iprover intro: ccontr notE ) |
15411 | 570 |
|
571 |
lemma contrapos_pp: |
|
60759 | 572 |
assumes p1: Q |
63575 | 573 |
and p2: "\<not> P \<Longrightarrow> \<not> Q" |
60759 | 574 |
shows P |
63575 | 575 |
by (iprover intro: classical p1 p2 notE) |
15411 | 576 |
|
577 |
||
60758 | 578 |
subsubsection \<open>Unique existence\<close> |
15411 | 579 |
|
71827 | 580 |
lemma Uniq_I [intro?]: |
581 |
assumes "\<And>x y. \<lbrakk>P x; P y\<rbrakk> \<Longrightarrow> y = x" |
|
582 |
shows "Uniq P" |
|
583 |
unfolding Uniq_def by (iprover intro: assms allI impI) |
|
584 |
||
585 |
lemma Uniq_D [dest?]: "\<lbrakk>Uniq P; P a; P b\<rbrakk> \<Longrightarrow> a=b" |
|
586 |
unfolding Uniq_def by (iprover dest: spec mp) |
|
587 |
||
15411 | 588 |
lemma ex1I: |
60759 | 589 |
assumes "P a" "\<And>x. P x \<Longrightarrow> x = a" |
590 |
shows "\<exists>!x. P x" |
|
63575 | 591 |
unfolding Ex1_def by (iprover intro: assms exI conjI allI impI) |
15411 | 592 |
|
63575 | 593 |
text \<open>Sometimes easier to use: the premises have no shared variables. Safe!\<close> |
15411 | 594 |
lemma ex_ex1I: |
60759 | 595 |
assumes ex_prem: "\<exists>x. P x" |
63575 | 596 |
and eq: "\<And>x y. \<lbrakk>P x; P y\<rbrakk> \<Longrightarrow> x = y" |
60759 | 597 |
shows "\<exists>!x. P x" |
63575 | 598 |
by (iprover intro: ex_prem [THEN exE] ex1I eq) |
15411 | 599 |
|
600 |
lemma ex1E: |
|
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
601 |
assumes major: "\<exists>!x. P x" and minor: "\<And>x. \<lbrakk>P x; \<forall>y. P y \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R" |
60759 | 602 |
shows R |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
603 |
proof (rule major [unfolded Ex1_def, THEN exE]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
604 |
show "\<And>x. P x \<and> (\<forall>y. P y \<longrightarrow> y = x) \<Longrightarrow> R" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
605 |
by (iprover intro: minor elim: conjE) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
606 |
qed |
15411 | 607 |
|
60759 | 608 |
lemma ex1_implies_ex: "\<exists>!x. P x \<Longrightarrow> \<exists>x. P x" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
609 |
by (iprover intro: exI elim: ex1E) |
15411 | 610 |
|
60758 | 611 |
subsubsection \<open>Classical intro rules for disjunction and existential quantifiers\<close> |
15411 | 612 |
|
613 |
lemma disjCI: |
|
63575 | 614 |
assumes "\<not> Q \<Longrightarrow> P" |
615 |
shows "P \<or> Q" |
|
616 |
by (rule classical) (iprover intro: assms disjI1 disjI2 notI elim: notE) |
|
15411 | 617 |
|
60759 | 618 |
lemma excluded_middle: "\<not> P \<or> P" |
63575 | 619 |
by (iprover intro: disjCI) |
15411 | 620 |
|
60758 | 621 |
text \<open> |
20944 | 622 |
case distinction as a natural deduction rule. |
63575 | 623 |
Note that \<open>\<not> P\<close> is the second case, not the first. |
60758 | 624 |
\<close> |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
625 |
lemma case_split [case_names True False]: |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
626 |
assumes "P \<Longrightarrow> Q" "\<not> P \<Longrightarrow> Q" |
60759 | 627 |
shows Q |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
628 |
using excluded_middle [of P] |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
629 |
by (iprover intro: assms elim: disjE) |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
630 |
|
63575 | 631 |
text \<open>Classical implies (\<open>\<longrightarrow>\<close>) elimination.\<close> |
15411 | 632 |
lemma impCE: |
60759 | 633 |
assumes major: "P \<longrightarrow> Q" |
63575 | 634 |
and minor: "\<not> P \<Longrightarrow> R" "Q \<Longrightarrow> R" |
60759 | 635 |
shows R |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
636 |
using excluded_middle [of P] |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
637 |
by (iprover intro: minor major [THEN mp] elim: disjE)+ |
15411 | 638 |
|
63575 | 639 |
text \<open> |
640 |
This version of \<open>\<longrightarrow>\<close> elimination works on \<open>Q\<close> before \<open>P\<close>. It works best for |
|
641 |
those cases in which \<open>P\<close> holds "almost everywhere". Can't install as |
|
642 |
default: would break old proofs. |
|
643 |
\<close> |
|
15411 | 644 |
lemma impCE': |
60759 | 645 |
assumes major: "P \<longrightarrow> Q" |
63575 | 646 |
and minor: "Q \<Longrightarrow> R" "\<not> P \<Longrightarrow> R" |
60759 | 647 |
shows R |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
648 |
using assms by (elim impCE) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
649 |
|
80948
572970d15ab0
To tiny but maybe useful lemmas (moved in from the AFP, Word_Lib)
paulson <lp15@cam.ac.uk>
parents:
80934
diff
changeset
|
650 |
text \<open>The analogous introduction rule for conjunction, above, is even constructive\<close> |
572970d15ab0
To tiny but maybe useful lemmas (moved in from the AFP, Word_Lib)
paulson <lp15@cam.ac.uk>
parents:
80934
diff
changeset
|
651 |
lemma context_disjE: |
572970d15ab0
To tiny but maybe useful lemmas (moved in from the AFP, Word_Lib)
paulson <lp15@cam.ac.uk>
parents:
80934
diff
changeset
|
652 |
assumes major: "P \<or> Q" and minor: "P \<Longrightarrow> R" "\<not>P \<Longrightarrow> Q \<Longrightarrow> R" |
572970d15ab0
To tiny but maybe useful lemmas (moved in from the AFP, Word_Lib)
paulson <lp15@cam.ac.uk>
parents:
80934
diff
changeset
|
653 |
shows R |
572970d15ab0
To tiny but maybe useful lemmas (moved in from the AFP, Word_Lib)
paulson <lp15@cam.ac.uk>
parents:
80934
diff
changeset
|
654 |
by (iprover intro: disjE [OF major] disjE [OF excluded_middle] assms) |
15411 | 655 |
|
63575 | 656 |
text \<open>Classical \<open>\<longleftrightarrow>\<close> elimination.\<close> |
15411 | 657 |
lemma iffCE: |
60759 | 658 |
assumes major: "P = Q" |
63575 | 659 |
and minor: "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R" "\<lbrakk>\<not> P; \<not> Q\<rbrakk> \<Longrightarrow> R" |
60759 | 660 |
shows R |
63575 | 661 |
by (rule major [THEN iffE]) (iprover intro: minor elim: impCE notE) |
15411 | 662 |
|
663 |
lemma exCI: |
|
60759 | 664 |
assumes "\<forall>x. \<not> P x \<Longrightarrow> P a" |
665 |
shows "\<exists>x. P x" |
|
63575 | 666 |
by (rule ccontr) (iprover intro: assms exI allI notI notE [of "\<exists>x. P x"]) |
15411 | 667 |
|
668 |
||
60758 | 669 |
subsubsection \<open>Intuitionistic Reasoning\<close> |
12386 | 670 |
|
671 |
lemma impE': |
|
60759 | 672 |
assumes 1: "P \<longrightarrow> Q" |
673 |
and 2: "Q \<Longrightarrow> R" |
|
674 |
and 3: "P \<longrightarrow> Q \<Longrightarrow> P" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
675 |
shows R |
12386 | 676 |
proof - |
677 |
from 3 and 1 have P . |
|
678 |
with 1 have Q by (rule impE) |
|
679 |
with 2 show R . |
|
680 |
qed |
|
681 |
||
682 |
lemma allE': |
|
60759 | 683 |
assumes 1: "\<forall>x. P x" |
684 |
and 2: "P x \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
685 |
shows Q |
12386 | 686 |
proof - |
687 |
from 1 have "P x" by (rule spec) |
|
688 |
from this and 1 show Q by (rule 2) |
|
689 |
qed |
|
690 |
||
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
691 |
lemma notE': |
60759 | 692 |
assumes 1: "\<not> P" |
693 |
and 2: "\<not> P \<Longrightarrow> P" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
694 |
shows R |
12386 | 695 |
proof - |
696 |
from 2 and 1 have P . |
|
697 |
with 1 show R by (rule notE) |
|
698 |
qed |
|
699 |
||
60759 | 700 |
lemma TrueE: "True \<Longrightarrow> P \<Longrightarrow> P" . |
701 |
lemma notFalseE: "\<not> False \<Longrightarrow> P \<Longrightarrow> P" . |
|
22444
fb80fedd192d
added safe intro rules for removing "True" subgoals as well as "~ False" ones.
dixon
parents:
22377
diff
changeset
|
702 |
|
22467
c9357ef01168
TrueElim and notTrueElim tested and added as safe elim rules.
dixon
parents:
22445
diff
changeset
|
703 |
lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE |
15801 | 704 |
and [Pure.intro!] = iffI conjI impI TrueI notI allI refl |
705 |
and [Pure.elim 2] = allE notE' impE' |
|
706 |
and [Pure.intro] = exI disjI2 disjI1 |
|
12386 | 707 |
|
708 |
lemmas [trans] = trans |
|
709 |
and [sym] = sym not_sym |
|
15801 | 710 |
and [Pure.elim?] = iffD1 iffD2 impE |
11750 | 711 |
|
11438
3d9222b80989
declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents:
11432
diff
changeset
|
712 |
|
60758 | 713 |
subsubsection \<open>Atomizing meta-level connectives\<close> |
11750 | 714 |
|
28513 | 715 |
axiomatization where |
63575 | 716 |
eq_reflection: "x = y \<Longrightarrow> x \<equiv> y" \<comment> \<open>admissible axiom\<close> |
28513 | 717 |
|
60759 | 718 |
lemma atomize_all [atomize]: "(\<And>x. P x) \<equiv> Trueprop (\<forall>x. P x)" |
12003 | 719 |
proof |
60759 | 720 |
assume "\<And>x. P x" |
721 |
then show "\<forall>x. P x" .. |
|
9488 | 722 |
next |
60759 | 723 |
assume "\<forall>x. P x" |
724 |
then show "\<And>x. P x" by (rule allE) |
|
9488 | 725 |
qed |
726 |
||
60759 | 727 |
lemma atomize_imp [atomize]: "(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)" |
12003 | 728 |
proof |
60759 | 729 |
assume r: "A \<Longrightarrow> B" |
730 |
show "A \<longrightarrow> B" by (rule impI) (rule r) |
|
9488 | 731 |
next |
60759 | 732 |
assume "A \<longrightarrow> B" and A |
23553 | 733 |
then show B by (rule mp) |
9488 | 734 |
qed |
735 |
||
60759 | 736 |
lemma atomize_not: "(A \<Longrightarrow> False) \<equiv> Trueprop (\<not> A)" |
14749 | 737 |
proof |
60759 | 738 |
assume r: "A \<Longrightarrow> False" |
739 |
show "\<not> A" by (rule notI) (rule r) |
|
14749 | 740 |
next |
60759 | 741 |
assume "\<not> A" and A |
23553 | 742 |
then show False by (rule notE) |
14749 | 743 |
qed |
744 |
||
60759 | 745 |
lemma atomize_eq [atomize, code]: "(x \<equiv> y) \<equiv> Trueprop (x = y)" |
12003 | 746 |
proof |
60759 | 747 |
assume "x \<equiv> y" |
748 |
show "x = y" by (unfold \<open>x \<equiv> y\<close>) (rule refl) |
|
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
749 |
next |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
750 |
assume "x = y" |
60759 | 751 |
then show "x \<equiv> y" by (rule eq_reflection) |
10432
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
752 |
qed |
3dfbc913d184
added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents:
10383
diff
changeset
|
753 |
|
60759 | 754 |
lemma atomize_conj [atomize]: "(A &&& B) \<equiv> Trueprop (A \<and> B)" |
12003 | 755 |
proof |
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28741
diff
changeset
|
756 |
assume conj: "A &&& B" |
60759 | 757 |
show "A \<and> B" |
19121 | 758 |
proof (rule conjI) |
759 |
from conj show A by (rule conjunctionD1) |
|
760 |
from conj show B by (rule conjunctionD2) |
|
761 |
qed |
|
11953 | 762 |
next |
60759 | 763 |
assume conj: "A \<and> B" |
28856
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents:
28741
diff
changeset
|
764 |
show "A &&& B" |
19121 | 765 |
proof - |
766 |
from conj show A .. |
|
767 |
from conj show B .. |
|
11953 | 768 |
qed |
769 |
qed |
|
770 |
||
12386 | 771 |
lemmas [symmetric, rulify] = atomize_all atomize_imp |
18832 | 772 |
and [symmetric, defn] = atomize_all atomize_imp atomize_eq |
12386 | 773 |
|
11750 | 774 |
|
60758 | 775 |
subsubsection \<open>Atomizing elimination rules\<close> |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
776 |
|
60759 | 777 |
lemma atomize_exL[atomize_elim]: "(\<And>x. P x \<Longrightarrow> Q) \<equiv> ((\<exists>x. P x) \<Longrightarrow> Q)" |
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74741
diff
changeset
|
778 |
by (rule equal_intr_rule) iprover+ |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
779 |
|
60759 | 780 |
lemma atomize_conjL[atomize_elim]: "(A \<Longrightarrow> B \<Longrightarrow> C) \<equiv> (A \<and> B \<Longrightarrow> C)" |
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74741
diff
changeset
|
781 |
by (rule equal_intr_rule) iprover+ |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
782 |
|
60759 | 783 |
lemma atomize_disjL[atomize_elim]: "((A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C) \<equiv> ((A \<or> B \<Longrightarrow> C) \<Longrightarrow> C)" |
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74741
diff
changeset
|
784 |
by (rule equal_intr_rule) iprover+ |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
785 |
|
60759 | 786 |
lemma atomize_elimL[atomize_elim]: "(\<And>B. (A \<Longrightarrow> B) \<Longrightarrow> B) \<equiv> Trueprop A" .. |
26580
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
787 |
|
c3e597a476fd
Generic conversion and tactic "atomize_elim" to convert elimination rules
krauss
parents:
26555
diff
changeset
|
788 |
|
60758 | 789 |
subsection \<open>Package setup\<close> |
20944 | 790 |
|
69605 | 791 |
ML_file \<open>Tools/hologic.ML\<close> |
70847 | 792 |
ML_file \<open>Tools/rewrite_hol_proof.ML\<close> |
51314
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51304
diff
changeset
|
793 |
|
70879 | 794 |
setup \<open>Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc Rewrite_HOL_Proof.rews)\<close> |
70849 | 795 |
|
51314
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51304
diff
changeset
|
796 |
|
60758 | 797 |
subsubsection \<open>Sledgehammer setup\<close> |
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35808
diff
changeset
|
798 |
|
60758 | 799 |
text \<open> |
63575 | 800 |
Theorems blacklisted to Sledgehammer. These theorems typically produce clauses |
801 |
that are prolific (match too many equality or membership literals) and relate to |
|
802 |
seldom-used facts. Some duplicate other rules. |
|
60758 | 803 |
\<close> |
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35808
diff
changeset
|
804 |
|
57963 | 805 |
named_theorems no_atp "theorems that should be filtered out by Sledgehammer" |
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35808
diff
changeset
|
806 |
|
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35808
diff
changeset
|
807 |
|
60758 | 808 |
subsubsection \<open>Classical Reasoner setup\<close> |
9529 | 809 |
|
60759 | 810 |
lemma imp_elim: "P \<longrightarrow> Q \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R" |
26411 | 811 |
by (rule classical) iprover |
812 |
||
60759 | 813 |
lemma swap: "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R" |
26411 | 814 |
by (rule classical) iprover |
815 |
||
62958
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62913
diff
changeset
|
816 |
lemma thin_refl: "\<lbrakk>x = x; PROP W\<rbrakk> \<Longrightarrow> PROP W" . |
20944 | 817 |
|
60758 | 818 |
ML \<open> |
42799 | 819 |
structure Hypsubst = Hypsubst |
820 |
( |
|
21218 | 821 |
val dest_eq = HOLogic.dest_eq |
21151 | 822 |
val dest_Trueprop = HOLogic.dest_Trueprop |
823 |
val dest_imp = HOLogic.dest_imp |
|
26411 | 824 |
val eq_reflection = @{thm eq_reflection} |
825 |
val rev_eq_reflection = @{thm meta_eq_to_obj_eq} |
|
826 |
val imp_intr = @{thm impI} |
|
827 |
val rev_mp = @{thm rev_mp} |
|
828 |
val subst = @{thm subst} |
|
829 |
val sym = @{thm sym} |
|
22129 | 830 |
val thin_refl = @{thm thin_refl}; |
42799 | 831 |
); |
21671 | 832 |
open Hypsubst; |
21151 | 833 |
|
42799 | 834 |
structure Classical = Classical |
835 |
( |
|
26411 | 836 |
val imp_elim = @{thm imp_elim} |
837 |
val not_elim = @{thm notE} |
|
838 |
val swap = @{thm swap} |
|
839 |
val classical = @{thm classical} |
|
21151 | 840 |
val sizef = Drule.size_of_thm |
841 |
val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] |
|
42799 | 842 |
); |
21151 | 843 |
|
58826 | 844 |
structure Basic_Classical: BASIC_CLASSICAL = Classical; |
33308
cf62d1690d04
separate ResBlacklist, based on scalable persistent data -- avoids inefficient hashing later on;
wenzelm
parents:
33185
diff
changeset
|
845 |
open Basic_Classical; |
60758 | 846 |
\<close> |
22129 | 847 |
|
60758 | 848 |
setup \<open> |
35389 | 849 |
(*prevent substitution on bool*) |
58826 | 850 |
let |
80662 | 851 |
fun non_bool_eq \<^Const_>\<open>HOL.eq T\<close> = T <> \<^Type>\<open>bool\<close> |
58826 | 852 |
| non_bool_eq _ = false; |
853 |
fun hyp_subst_tac' ctxt = |
|
854 |
SUBGOAL (fn (goal, i) => |
|
80662 | 855 |
if Term.exists_subterm non_bool_eq goal |
58826 | 856 |
then Hypsubst.hyp_subst_tac ctxt i |
857 |
else no_tac); |
|
858 |
in |
|
859 |
Context_Rules.addSWrapper (fn ctxt => fn tac => hyp_subst_tac' ctxt ORELSE' tac) |
|
860 |
end |
|
60758 | 861 |
\<close> |
21009 | 862 |
|
863 |
declare iffI [intro!] |
|
864 |
and notI [intro!] |
|
865 |
and impI [intro!] |
|
866 |
and disjCI [intro!] |
|
867 |
and conjI [intro!] |
|
868 |
and TrueI [intro!] |
|
869 |
and refl [intro!] |
|
870 |
||
871 |
declare iffCE [elim!] |
|
872 |
and FalseE [elim!] |
|
873 |
and impCE [elim!] |
|
874 |
and disjE [elim!] |
|
875 |
and conjE [elim!] |
|
876 |
||
877 |
declare ex_ex1I [intro!] |
|
878 |
and allI [intro!] |
|
879 |
and exI [intro] |
|
880 |
||
881 |
declare exE [elim!] |
|
882 |
allE [elim] |
|
883 |
||
69593 | 884 |
ML \<open>val HOL_cs = claset_of \<^context>\<close> |
19162 | 885 |
|
60759 | 886 |
lemma contrapos_np: "\<not> Q \<Longrightarrow> (\<not> P \<Longrightarrow> Q) \<Longrightarrow> P" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
887 |
by (erule swap) |
10383 | 888 |
|
18689
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
889 |
declare ex_ex1I [rule del, intro! 2] |
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
890 |
and ex1I [intro] |
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents:
18595
diff
changeset
|
891 |
|
41865
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41827
diff
changeset
|
892 |
declare ext [intro] |
4e8483cc2cc5
declare ext [intro]: Extensionality now available by default
paulson
parents:
41827
diff
changeset
|
893 |
|
12386 | 894 |
lemmas [intro?] = ext |
895 |
and [elim?] = ex1_implies_ex |
|
11977 | 896 |
|
63575 | 897 |
text \<open>Better than \<open>ex1E\<close> for classical reasoner: needs no quantifier duplication!\<close> |
20973 | 898 |
lemma alt_ex1E [elim!]: |
20944 | 899 |
assumes major: "\<exists>!x. P x" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
900 |
and minor: "\<And>x. \<lbrakk>P x; \<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'\<rbrakk> \<Longrightarrow> R" |
20944 | 901 |
shows R |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
902 |
proof (rule ex1E [OF major minor]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
903 |
show "\<forall>y y'. P y \<and> P y' \<longrightarrow> y = y'" if "P x" and \<section>: "\<forall>y. P y \<longrightarrow> y = x" for x |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
904 |
using \<open>P x\<close> \<section> \<section> by fast |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
905 |
qed assumption |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
906 |
|
71827 | 907 |
text \<open>And again using Uniq\<close> |
908 |
lemma alt_ex1E': |
|
909 |
assumes "\<exists>!x. P x" "\<And>x. \<lbrakk>P x; \<exists>\<^sub>\<le>\<^sub>1x. P x\<rbrakk> \<Longrightarrow> R" |
|
910 |
shows R |
|
911 |
using assms unfolding Uniq_def by fast |
|
912 |
||
913 |
lemma ex1_iff_ex_Uniq: "(\<exists>!x. P x) \<longleftrightarrow> (\<exists>x. P x) \<and> (\<exists>\<^sub>\<le>\<^sub>1x. P x)" |
|
914 |
unfolding Uniq_def by fast |
|
915 |
||
20944 | 916 |
|
60758 | 917 |
ML \<open> |
42477 | 918 |
structure Blast = Blast |
919 |
( |
|
920 |
structure Classical = Classical |
|
74383 | 921 |
val Trueprop_const = dest_Const \<^Const>\<open>Trueprop\<close> |
69593 | 922 |
val equality_name = \<^const_name>\<open>HOL.eq\<close> |
923 |
val not_name = \<^const_name>\<open>Not\<close> |
|
42477 | 924 |
val notE = @{thm notE} |
925 |
val ccontr = @{thm ccontr} |
|
926 |
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac |
|
927 |
); |
|
928 |
val blast_tac = Blast.blast_tac; |
|
60758 | 929 |
\<close> |
20944 | 930 |
|
931 |
||
60758 | 932 |
subsubsection \<open>THE: definite description operator\<close> |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
933 |
|
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
934 |
lemma the_equality [intro]: |
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
935 |
assumes "P a" |
63575 | 936 |
and "\<And>x. P x \<Longrightarrow> x = a" |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
937 |
shows "(THE x. P x) = a" |
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
938 |
by (blast intro: assms trans [OF arg_cong [where f=The] the_eq_trivial]) |
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
939 |
|
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
940 |
lemma theI: |
63575 | 941 |
assumes "P a" |
942 |
and "\<And>x. P x \<Longrightarrow> x = a" |
|
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
943 |
shows "P (THE x. P x)" |
63575 | 944 |
by (iprover intro: assms the_equality [THEN ssubst]) |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
945 |
|
60759 | 946 |
lemma theI': "\<exists>!x. P x \<Longrightarrow> P (THE x. P x)" |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
947 |
by (blast intro: theI) |
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
948 |
|
63575 | 949 |
text \<open>Easier to apply than \<open>theI\<close>: only one occurrence of \<open>P\<close>.\<close> |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
950 |
lemma theI2: |
60759 | 951 |
assumes "P a" "\<And>x. P x \<Longrightarrow> x = a" "\<And>x. P x \<Longrightarrow> Q x" |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
952 |
shows "Q (THE x. P x)" |
63575 | 953 |
by (iprover intro: assms theI) |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
954 |
|
63575 | 955 |
lemma the1I2: |
956 |
assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x" |
|
957 |
shows "Q (THE x. P x)" |
|
958 |
by (iprover intro: assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)] elim: allE impE) |
|
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
959 |
|
60759 | 960 |
lemma the1_equality [elim?]: "\<lbrakk>\<exists>!x. P x; P a\<rbrakk> \<Longrightarrow> (THE x. P x) = a" |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
961 |
by blast |
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
962 |
|
71827 | 963 |
lemma the1_equality': "\<lbrakk>\<exists>\<^sub>\<le>\<^sub>1x. P x; P a\<rbrakk> \<Longrightarrow> (THE x. P x) = a" |
964 |
unfolding Uniq_def by blast |
|
965 |
||
60759 | 966 |
lemma the_sym_eq_trivial: "(THE y. x = y) = x" |
59504
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
967 |
by blast |
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
968 |
|
8c6747dba731
New lemmas and a bit of tidying up.
paulson <lp15@cam.ac.uk>
parents:
59028
diff
changeset
|
969 |
|
60758 | 970 |
subsubsection \<open>Simplifier\<close> |
12281 | 971 |
|
60759 | 972 |
lemma eta_contract_eq: "(\<lambda>s. f s) = f" .. |
12281 | 973 |
|
71918 | 974 |
lemma subst_all: |
975 |
\<open>(\<And>x. x = a \<Longrightarrow> PROP P x) \<equiv> PROP P a\<close> |
|
976 |
\<open>(\<And>x. a = x \<Longrightarrow> PROP P x) \<equiv> PROP P a\<close> |
|
71959 | 977 |
proof - |
978 |
show \<open>(\<And>x. x = a \<Longrightarrow> PROP P x) \<equiv> PROP P a\<close> |
|
979 |
proof (rule equal_intr_rule) |
|
980 |
assume *: \<open>\<And>x. x = a \<Longrightarrow> PROP P x\<close> |
|
981 |
show \<open>PROP P a\<close> |
|
982 |
by (rule *) (rule refl) |
|
983 |
next |
|
984 |
fix x |
|
985 |
assume \<open>PROP P a\<close> and \<open>x = a\<close> |
|
986 |
from \<open>x = a\<close> have \<open>x \<equiv> a\<close> |
|
987 |
by (rule eq_reflection) |
|
988 |
with \<open>PROP P a\<close> show \<open>PROP P x\<close> |
|
989 |
by simp |
|
990 |
qed |
|
991 |
show \<open>(\<And>x. a = x \<Longrightarrow> PROP P x) \<equiv> PROP P a\<close> |
|
992 |
proof (rule equal_intr_rule) |
|
993 |
assume *: \<open>\<And>x. a = x \<Longrightarrow> PROP P x\<close> |
|
994 |
show \<open>PROP P a\<close> |
|
995 |
by (rule *) (rule refl) |
|
996 |
next |
|
997 |
fix x |
|
998 |
assume \<open>PROP P a\<close> and \<open>a = x\<close> |
|
999 |
from \<open>a = x\<close> have \<open>a \<equiv> x\<close> |
|
1000 |
by (rule eq_reflection) |
|
1001 |
with \<open>PROP P a\<close> show \<open>PROP P x\<close> |
|
1002 |
by simp |
|
1003 |
qed |
|
71918 | 1004 |
qed |
1005 |
||
12281 | 1006 |
lemma simp_thms: |
60759 | 1007 |
shows not_not: "(\<not> \<not> P) = P" |
1008 |
and Not_eq_iff: "((\<not> P) = (\<not> Q)) = (P = Q)" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12892
diff
changeset
|
1009 |
and |
60759 | 1010 |
"(P \<noteq> Q) = (P = (\<not> Q))" |
1011 |
"(P \<or> \<not>P) = True" "(\<not> P \<or> P) = True" |
|
12281 | 1012 |
"(x = x) = True" |
32068 | 1013 |
and not_True_eq_False [code]: "(\<not> True) = False" |
1014 |
and not_False_eq_True [code]: "(\<not> False) = True" |
|
20944 | 1015 |
and |
60759 | 1016 |
"(\<not> P) \<noteq> P" "P \<noteq> (\<not> P)" |
1017 |
"(True = P) = P" |
|
20944 | 1018 |
and eq_True: "(P = True) = P" |
60759 | 1019 |
and "(False = P) = (\<not> P)" |
20944 | 1020 |
and eq_False: "(P = False) = (\<not> P)" |
1021 |
and |
|
60759 | 1022 |
"(True \<longrightarrow> P) = P" "(False \<longrightarrow> P) = True" |
1023 |
"(P \<longrightarrow> True) = True" "(P \<longrightarrow> P) = True" |
|
1024 |
"(P \<longrightarrow> False) = (\<not> P)" "(P \<longrightarrow> \<not> P) = (\<not> P)" |
|
1025 |
"(P \<and> True) = P" "(True \<and> P) = P" |
|
1026 |
"(P \<and> False) = False" "(False \<and> P) = False" |
|
1027 |
"(P \<and> P) = P" "(P \<and> (P \<and> Q)) = (P \<and> Q)" |
|
1028 |
"(P \<and> \<not> P) = False" "(\<not> P \<and> P) = False" |
|
1029 |
"(P \<or> True) = True" "(True \<or> P) = True" |
|
1030 |
"(P \<or> False) = P" "(False \<or> P) = P" |
|
1031 |
"(P \<or> P) = P" "(P \<or> (P \<or> Q)) = (P \<or> Q)" and |
|
1032 |
"(\<forall>x. P) = P" "(\<exists>x. P) = P" "\<exists>x. x = t" "\<exists>x. t = x" |
|
31166
a90fe83f58ea
"{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents:
31156
diff
changeset
|
1033 |
and |
60759 | 1034 |
"\<And>P. (\<exists>x. x = t \<and> P x) = P t" |
1035 |
"\<And>P. (\<exists>x. t = x \<and> P x) = P t" |
|
1036 |
"\<And>P. (\<forall>x. x = t \<longrightarrow> P x) = P t" |
|
1037 |
"\<And>P. (\<forall>x. t = x \<longrightarrow> P x) = P t" |
|
66109 | 1038 |
"(\<forall>x. x \<noteq> t) = False" "(\<forall>x. t \<noteq> x) = False" |
17589 | 1039 |
by (blast, blast, blast, blast, blast, iprover+) |
13421 | 1040 |
|
63575 | 1041 |
lemma disj_absorb: "A \<or> A \<longleftrightarrow> A" |
14201 | 1042 |
by blast |
1043 |
||
63575 | 1044 |
lemma disj_left_absorb: "A \<or> (A \<or> B) \<longleftrightarrow> A \<or> B" |
14201 | 1045 |
by blast |
1046 |
||
63575 | 1047 |
lemma conj_absorb: "A \<and> A \<longleftrightarrow> A" |
14201 | 1048 |
by blast |
1049 |
||
63575 | 1050 |
lemma conj_left_absorb: "A \<and> (A \<and> B) \<longleftrightarrow> A \<and> B" |
14201 | 1051 |
by blast |
1052 |
||
12281 | 1053 |
lemma eq_ac: |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56941
diff
changeset
|
1054 |
shows eq_commute: "a = b \<longleftrightarrow> b = a" |
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56941
diff
changeset
|
1055 |
and iff_left_commute: "(P \<longleftrightarrow> (Q \<longleftrightarrow> R)) \<longleftrightarrow> (Q \<longleftrightarrow> (P \<longleftrightarrow> R))" |
63575 | 1056 |
and iff_assoc: "((P \<longleftrightarrow> Q) \<longleftrightarrow> R) \<longleftrightarrow> (P \<longleftrightarrow> (Q \<longleftrightarrow> R))" |
1057 |
by (iprover, blast+) |
|
1058 |
||
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
56941
diff
changeset
|
1059 |
lemma neq_commute: "a \<noteq> b \<longleftrightarrow> b \<noteq> a" by iprover |
12281 | 1060 |
|
1061 |
lemma conj_comms: |
|
63575 | 1062 |
shows conj_commute: "P \<and> Q \<longleftrightarrow> Q \<and> P" |
1063 |
and conj_left_commute: "P \<and> (Q \<and> R) \<longleftrightarrow> Q \<and> (P \<and> R)" by iprover+ |
|
1064 |
lemma conj_assoc: "(P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)" by iprover |
|
12281 | 1065 |
|
19174 | 1066 |
lemmas conj_ac = conj_commute conj_left_commute conj_assoc |
1067 |
||
12281 | 1068 |
lemma disj_comms: |
63575 | 1069 |
shows disj_commute: "P \<or> Q \<longleftrightarrow> Q \<or> P" |
1070 |
and disj_left_commute: "P \<or> (Q \<or> R) \<longleftrightarrow> Q \<or> (P \<or> R)" by iprover+ |
|
1071 |
lemma disj_assoc: "(P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)" by iprover |
|
12281 | 1072 |
|
19174 | 1073 |
lemmas disj_ac = disj_commute disj_left_commute disj_assoc |
1074 |
||
63575 | 1075 |
lemma conj_disj_distribL: "P \<and> (Q \<or> R) \<longleftrightarrow> P \<and> Q \<or> P \<and> R" by iprover |
1076 |
lemma conj_disj_distribR: "(P \<or> Q) \<and> R \<longleftrightarrow> P \<and> R \<or> Q \<and> R" by iprover |
|
12281 | 1077 |
|
63575 | 1078 |
lemma disj_conj_distribL: "P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)" by iprover |
1079 |
lemma disj_conj_distribR: "(P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)" by iprover |
|
12281 | 1080 |
|
60759 | 1081 |
lemma imp_conjR: "(P \<longrightarrow> (Q \<and> R)) = ((P \<longrightarrow> Q) \<and> (P \<longrightarrow> R))" by iprover |
1082 |
lemma imp_conjL: "((P \<and> Q) \<longrightarrow> R) = (P \<longrightarrow> (Q \<longrightarrow> R))" by iprover |
|
1083 |
lemma imp_disjL: "((P \<or> Q) \<longrightarrow> R) = ((P \<longrightarrow> R) \<and> (Q \<longrightarrow> R))" by iprover |
|
12281 | 1084 |
|
61799 | 1085 |
text \<open>These two are specialized, but \<open>imp_disj_not1\<close> is useful in \<open>Auth/Yahalom\<close>.\<close> |
63575 | 1086 |
lemma imp_disj_not1: "(P \<longrightarrow> Q \<or> R) \<longleftrightarrow> (\<not> Q \<longrightarrow> P \<longrightarrow> R)" by blast |
1087 |
lemma imp_disj_not2: "(P \<longrightarrow> Q \<or> R) \<longleftrightarrow> (\<not> R \<longrightarrow> P \<longrightarrow> Q)" by blast |
|
12281 | 1088 |
|
63575 | 1089 |
lemma imp_disj1: "((P \<longrightarrow> Q) \<or> R) \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)" by blast |
1090 |
lemma imp_disj2: "(Q \<or> (P \<longrightarrow> R)) \<longleftrightarrow> (P \<longrightarrow> Q \<or> R)" by blast |
|
12281 | 1091 |
|
63575 | 1092 |
lemma imp_cong: "(P = P') \<Longrightarrow> (P' \<Longrightarrow> (Q = Q')) \<Longrightarrow> ((P \<longrightarrow> Q) \<longleftrightarrow> (P' \<longrightarrow> Q'))" |
21151 | 1093 |
by iprover |
1094 |
||
63575 | 1095 |
lemma de_Morgan_disj: "\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q" by iprover |
1096 |
lemma de_Morgan_conj: "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q" by blast |
|
1097 |
lemma not_imp: "\<not> (P \<longrightarrow> Q) \<longleftrightarrow> P \<and> \<not> Q" by blast |
|
1098 |
lemma not_iff: "P \<noteq> Q \<longleftrightarrow> (P \<longleftrightarrow> \<not> Q)" by blast |
|
1099 |
lemma disj_not1: "\<not> P \<or> Q \<longleftrightarrow> (P \<longrightarrow> Q)" by blast |
|
1100 |
lemma disj_not2: "P \<or> \<not> Q \<longleftrightarrow> (Q \<longrightarrow> P)" by blast \<comment> \<open>changes orientation :-(\<close> |
|
1101 |
lemma imp_conv_disj: "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P) \<or> Q" by blast |
|
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
62958
diff
changeset
|
1102 |
lemma disj_imp: "P \<or> Q \<longleftrightarrow> \<not> P \<longrightarrow> Q" by blast |
12281 | 1103 |
|
63575 | 1104 |
lemma iff_conv_conj_imp: "(P \<longleftrightarrow> Q) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)" by iprover |
12281 | 1105 |
|
1106 |
||
63575 | 1107 |
lemma cases_simp: "(P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> Q) \<longleftrightarrow> Q" |
62390 | 1108 |
\<comment> \<open>Avoids duplication of subgoals after \<open>if_split\<close>, when the true and false\<close> |
61799 | 1109 |
\<comment> \<open>cases boil down to the same thing.\<close> |
12281 | 1110 |
by blast |
1111 |
||
63575 | 1112 |
lemma not_all: "\<not> (\<forall>x. P x) \<longleftrightarrow> (\<exists>x. \<not> P x)" by blast |
1113 |
lemma imp_all: "((\<forall>x. P x) \<longrightarrow> Q) \<longleftrightarrow> (\<exists>x. P x \<longrightarrow> Q)" by blast |
|
1114 |
lemma not_ex: "\<not> (\<exists>x. P x) \<longleftrightarrow> (\<forall>x. \<not> P x)" by iprover |
|
1115 |
lemma imp_ex: "((\<exists>x. P x) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P x \<longrightarrow> Q)" by iprover |
|
1116 |
lemma all_not_ex: "(\<forall>x. P x) \<longleftrightarrow> \<not> (\<exists>x. \<not> P x)" by blast |
|
12281 | 1117 |
|
35828
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
blanchet
parents:
35808
diff
changeset
|
1118 |
declare All_def [no_atp] |
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24280
diff
changeset
|
1119 |
|
63575 | 1120 |
lemma ex_disj_distrib: "(\<exists>x. P x \<or> Q x) \<longleftrightarrow> (\<exists>x. P x) \<or> (\<exists>x. Q x)" by iprover |
1121 |
lemma all_conj_distrib: "(\<forall>x. P x \<and> Q x) \<longleftrightarrow> (\<forall>x. P x) \<and> (\<forall>x. Q x)" by iprover |
|
79772
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
78800
diff
changeset
|
1122 |
lemma all_imp_conj_distrib: "(\<forall>x. P x \<longrightarrow> Q x \<and> R x) \<longleftrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<and> (\<forall>x. P x \<longrightarrow> R x)" |
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
paulson <lp15@cam.ac.uk>
parents:
78800
diff
changeset
|
1123 |
by iprover |
12281 | 1124 |
|
60758 | 1125 |
text \<open> |
63575 | 1126 |
\<^medskip> The \<open>\<and>\<close> congruence rule: not included by default! |
60758 | 1127 |
May slow rewrite proofs down by as much as 50\%\<close> |
12281 | 1128 |
|
63575 | 1129 |
lemma conj_cong: "P = P' \<Longrightarrow> (P' \<Longrightarrow> Q = Q') \<Longrightarrow> (P \<and> Q) = (P' \<and> Q')" |
17589 | 1130 |
by iprover |
12281 | 1131 |
|
63575 | 1132 |
lemma rev_conj_cong: "Q = Q' \<Longrightarrow> (Q' \<Longrightarrow> P = P') \<Longrightarrow> (P \<and> Q) = (P' \<and> Q')" |
17589 | 1133 |
by iprover |
12281 | 1134 |
|
61799 | 1135 |
text \<open>The \<open>|\<close> congruence rule: not included by default!\<close> |
12281 | 1136 |
|
63575 | 1137 |
lemma disj_cong: "P = P' \<Longrightarrow> (\<not> P' \<Longrightarrow> Q = Q') \<Longrightarrow> (P \<or> Q) = (P' \<or> Q')" |
12281 | 1138 |
by blast |
1139 |
||
1140 |
||
63575 | 1141 |
text \<open>\<^medskip> if-then-else rules\<close> |
12281 | 1142 |
|
32068 | 1143 |
lemma if_True [code]: "(if True then x else y) = x" |
63575 | 1144 |
unfolding If_def by blast |
12281 | 1145 |
|
32068 | 1146 |
lemma if_False [code]: "(if False then x else y) = y" |
63575 | 1147 |
unfolding If_def by blast |
12281 | 1148 |
|
60759 | 1149 |
lemma if_P: "P \<Longrightarrow> (if P then x else y) = x" |
63575 | 1150 |
unfolding If_def by blast |
12281 | 1151 |
|
60759 | 1152 |
lemma if_not_P: "\<not> P \<Longrightarrow> (if P then x else y) = y" |
63575 | 1153 |
unfolding If_def by blast |
12281 | 1154 |
|
62390 | 1155 |
lemma if_split: "P (if Q then x else y) = ((Q \<longrightarrow> P x) \<and> (\<not> Q \<longrightarrow> P y))" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1156 |
proof (rule case_split [of Q]) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1157 |
show ?thesis if Q |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1158 |
using that by (simplesubst if_P) blast+ |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1159 |
show ?thesis if "\<not> Q" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1160 |
using that by (simplesubst if_not_P) blast+ |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1161 |
qed |
12281 | 1162 |
|
62390 | 1163 |
lemma if_split_asm: "P (if Q then x else y) = (\<not> ((Q \<and> \<not> P x) \<or> (\<not> Q \<and> \<not> P y)))" |
63575 | 1164 |
by (simplesubst if_split) blast |
12281 | 1165 |
|
62390 | 1166 |
lemmas if_splits [no_atp] = if_split if_split_asm |
12281 | 1167 |
|
1168 |
lemma if_cancel: "(if c then x else x) = x" |
|
63575 | 1169 |
by (simplesubst if_split) blast |
12281 | 1170 |
|
1171 |
lemma if_eq_cancel: "(if x = y then y else x) = x" |
|
63575 | 1172 |
by (simplesubst if_split) blast |
12281 | 1173 |
|
60759 | 1174 |
lemma if_bool_eq_conj: "(if P then Q else R) = ((P \<longrightarrow> Q) \<and> (\<not> P \<longrightarrow> R))" |
61799 | 1175 |
\<comment> \<open>This form is useful for expanding \<open>if\<close>s on the RIGHT of the \<open>\<Longrightarrow>\<close> symbol.\<close> |
62390 | 1176 |
by (rule if_split) |
12281 | 1177 |
|
60759 | 1178 |
lemma if_bool_eq_disj: "(if P then Q else R) = ((P \<and> Q) \<or> (\<not> P \<and> R))" |
61799 | 1179 |
\<comment> \<open>And this form is useful for expanding \<open>if\<close>s on the LEFT.\<close> |
62390 | 1180 |
by (simplesubst if_split) blast |
12281 | 1181 |
|
63575 | 1182 |
lemma Eq_TrueI: "P \<Longrightarrow> P \<equiv> True" unfolding atomize_eq by iprover |
1183 |
lemma Eq_FalseI: "\<not> P \<Longrightarrow> P \<equiv> False" unfolding atomize_eq by iprover |
|
12281 | 1184 |
|
63575 | 1185 |
text \<open>\<^medskip> let rules for simproc\<close> |
15423 | 1186 |
|
60759 | 1187 |
lemma Let_folded: "f x \<equiv> g x \<Longrightarrow> Let x f \<equiv> Let x g" |
15423 | 1188 |
by (unfold Let_def) |
1189 |
||
60759 | 1190 |
lemma Let_unfold: "f x \<equiv> g \<Longrightarrow> Let x f \<equiv> g" |
15423 | 1191 |
by (unfold Let_def) |
1192 |
||
60758 | 1193 |
text \<open> |
16999 | 1194 |
The following copy of the implication operator is useful for |
1195 |
fine-tuning congruence rules. It instructs the simplifier to simplify |
|
1196 |
its premise. |
|
60758 | 1197 |
\<close> |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1198 |
|
80932
261cd8722677
standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents:
80762
diff
changeset
|
1199 |
definition simp_implies :: "prop \<Rightarrow> prop \<Rightarrow> prop" (infixr \<open>=simp=>\<close> 1) |
67399 | 1200 |
where "simp_implies \<equiv> (\<Longrightarrow>)" |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1201 |
|
18457 | 1202 |
lemma simp_impliesI: |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1203 |
assumes PQ: "(PROP P \<Longrightarrow> PROP Q)" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1204 |
shows "PROP P =simp=> PROP Q" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1205 |
unfolding simp_implies_def |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1206 |
by (iprover intro: PQ) |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1207 |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1208 |
lemma simp_impliesE: |
25388 | 1209 |
assumes PQ: "PROP P =simp=> PROP Q" |
63575 | 1210 |
and P: "PROP P" |
1211 |
and QR: "PROP Q \<Longrightarrow> PROP R" |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1212 |
shows "PROP R" |
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1213 |
by (iprover intro: QR P PQ [unfolded simp_implies_def]) |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1214 |
|
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1215 |
lemma simp_implies_cong: |
60759 | 1216 |
assumes PP' :"PROP P \<equiv> PROP P'" |
63575 | 1217 |
and P'QQ': "PROP P' \<Longrightarrow> (PROP Q \<equiv> PROP Q')" |
60759 | 1218 |
shows "(PROP P =simp=> PROP Q) \<equiv> (PROP P' =simp=> PROP Q')" |
63575 | 1219 |
unfolding simp_implies_def |
1220 |
proof (rule equal_intr_rule) |
|
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1221 |
assume PQ: "PROP P \<Longrightarrow> PROP Q" |
63575 | 1222 |
and P': "PROP P'" |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1223 |
from PP' [symmetric] and P' have "PROP P" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1224 |
by (rule equal_elim_rule1) |
23553 | 1225 |
then have "PROP Q" by (rule PQ) |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1226 |
with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1227 |
next |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1228 |
assume P'Q': "PROP P' \<Longrightarrow> PROP Q'" |
63575 | 1229 |
and P: "PROP P" |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1230 |
from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) |
23553 | 1231 |
then have "PROP Q'" by (rule P'Q') |
16633
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1232 |
with P'QQ' [OF P', symmetric] show "PROP Q" |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1233 |
by (rule equal_elim_rule1) |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1234 |
qed |
208ebc9311f2
Implemented trick (due to Tobias Nipkow) for fine-tuning simplification
berghofe
parents:
16587
diff
changeset
|
1235 |
|
20944 | 1236 |
lemma uncurry: |
1237 |
assumes "P \<longrightarrow> Q \<longrightarrow> R" |
|
1238 |
shows "P \<and> Q \<longrightarrow> R" |
|
23553 | 1239 |
using assms by blast |
20944 | 1240 |
|
1241 |
lemma iff_allI: |
|
1242 |
assumes "\<And>x. P x = Q x" |
|
1243 |
shows "(\<forall>x. P x) = (\<forall>x. Q x)" |
|
23553 | 1244 |
using assms by blast |
20944 | 1245 |
|
1246 |
lemma iff_exI: |
|
1247 |
assumes "\<And>x. P x = Q x" |
|
1248 |
shows "(\<exists>x. P x) = (\<exists>x. Q x)" |
|
23553 | 1249 |
using assms by blast |
20944 | 1250 |
|
63575 | 1251 |
lemma all_comm: "(\<forall>x y. P x y) = (\<forall>y x. P x y)" |
20944 | 1252 |
by blast |
1253 |
||
63575 | 1254 |
lemma ex_comm: "(\<exists>x y. P x y) = (\<exists>y x. P x y)" |
20944 | 1255 |
by blast |
1256 |
||
69605 | 1257 |
ML_file \<open>Tools/simpdata.ML\<close> |
60758 | 1258 |
ML \<open>open Simpdata\<close> |
42455 | 1259 |
|
60758 | 1260 |
setup \<open> |
58826 | 1261 |
map_theory_simpset (put_simpset HOL_basic_ss) #> |
1262 |
Simplifier.method_setup Splitter.split_modifiers |
|
60758 | 1263 |
\<close> |
42455 | 1264 |
|
71886 | 1265 |
simproc_setup defined_Ex ("\<exists>x. P x") = \<open>K Quantifier1.rearrange_Ex\<close> |
1266 |
simproc_setup defined_All ("\<forall>x. P x") = \<open>K Quantifier1.rearrange_All\<close> |
|
71914 | 1267 |
simproc_setup defined_all("\<And>x. PROP P x") = \<open>K Quantifier1.rearrange_all\<close> |
1268 |
||
61799 | 1269 |
text \<open>Simproc for proving \<open>(y = x) \<equiv> False\<close> from premise \<open>\<not> (x = y)\<close>:\<close> |
24035 | 1270 |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1271 |
simproc_setup neq ("x = y") = \<open> |
63575 | 1272 |
let |
1273 |
val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI}; |
|
1274 |
fun is_neq eq lhs rhs thm = |
|
1275 |
(case Thm.prop_of thm of |
|
1276 |
_ $ (Not $ (eq' $ l' $ r')) => |
|
1277 |
Not = HOLogic.Not andalso eq' = eq andalso |
|
1278 |
r' aconv lhs andalso l' aconv rhs |
|
1279 |
| _ => false); |
|
1280 |
fun proc ss ct = |
|
1281 |
(case Thm.term_of ct of |
|
1282 |
eq $ lhs $ rhs => |
|
1283 |
(case find_first (is_neq eq lhs rhs) (Simplifier.prems_of ss) of |
|
1284 |
SOME thm => SOME (thm RS neq_to_EQ_False) |
|
1285 |
| NONE => NONE) |
|
1286 |
| _ => NONE); |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1287 |
in K proc end |
60758 | 1288 |
\<close> |
24035 | 1289 |
|
60758 | 1290 |
simproc_setup let_simp ("Let x f") = \<open> |
63575 | 1291 |
let |
1292 |
fun count_loose (Bound i) k = if i >= k then 1 else 0 |
|
1293 |
| count_loose (s $ t) k = count_loose s k + count_loose t k |
|
1294 |
| count_loose (Abs (_, _, t)) k = count_loose t (k + 1) |
|
1295 |
| count_loose _ _ = 0; |
|
80662 | 1296 |
fun is_trivial_let \<^Const_>\<open>Let _ _ for x t\<close> = |
63575 | 1297 |
(case t of |
1298 |
Abs (_, _, t') => count_loose t' 0 <= 1 |
|
1299 |
| _ => true); |
|
1300 |
in |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1301 |
K (fn ctxt => fn ct => |
63575 | 1302 |
if is_trivial_let (Thm.term_of ct) |
1303 |
then SOME @{thm Let_def} (*no or one ocurrence of bound variable*) |
|
1304 |
else |
|
1305 |
let (*Norbert Schirmer's case*) |
|
1306 |
val t = Thm.term_of ct; |
|
70326 | 1307 |
val (t', ctxt') = yield_singleton (Variable.import_terms false) t ctxt; |
63575 | 1308 |
in |
1309 |
Option.map (hd o Variable.export ctxt' ctxt o single) |
|
80662 | 1310 |
(case t' of \<^Const_>\<open>Let _ _ for x f\<close> => (* x and f are already in normal form *) |
63575 | 1311 |
if is_Free x orelse is_Bound x orelse is_Const x |
1312 |
then SOME @{thm Let_def} |
|
1313 |
else |
|
1314 |
let |
|
1315 |
val n = case f of (Abs (x, _, _)) => x | _ => "x"; |
|
1316 |
val cx = Thm.cterm_of ctxt x; |
|
1317 |
val xT = Thm.typ_of_cterm cx; |
|
1318 |
val cf = Thm.cterm_of ctxt f; |
|
1319 |
val fx_g = Simplifier.rewrite ctxt (Thm.apply cf cx); |
|
1320 |
val (_ $ _ $ g) = Thm.prop_of fx_g; |
|
1321 |
val g' = abstract_over (x, g); |
|
1322 |
val abs_g'= Abs (n, xT, g'); |
|
1323 |
in |
|
1324 |
if g aconv g' then |
|
1325 |
let |
|
1326 |
val rl = |
|
1327 |
infer_instantiate ctxt [(("f", 0), cf), (("x", 0), cx)] @{thm Let_unfold}; |
|
1328 |
in SOME (rl OF [fx_g]) end |
|
1329 |
else if (Envir.beta_eta_contract f) aconv (Envir.beta_eta_contract abs_g') |
|
1330 |
then NONE (*avoid identity conversion*) |
|
1331 |
else |
|
1332 |
let |
|
1333 |
val g'x = abs_g' $ x; |
|
1334 |
val g_g'x = Thm.symmetric (Thm.beta_conversion false (Thm.cterm_of ctxt g'x)); |
|
1335 |
val rl = |
|
1336 |
@{thm Let_folded} |> infer_instantiate ctxt |
|
1337 |
[(("f", 0), Thm.cterm_of ctxt f), |
|
1338 |
(("x", 0), cx), |
|
1339 |
(("g", 0), Thm.cterm_of ctxt abs_g')]; |
|
1340 |
in SOME (rl OF [Thm.transitive fx_g g_g'x]) end |
|
1341 |
end |
|
1342 |
| _ => NONE) |
|
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1343 |
end) |
63575 | 1344 |
end |
1345 |
\<close> |
|
24035 | 1346 |
|
21151 | 1347 |
lemma True_implies_equals: "(True \<Longrightarrow> PROP P) \<equiv> PROP P" |
1348 |
proof |
|
23389 | 1349 |
assume "True \<Longrightarrow> PROP P" |
1350 |
from this [OF TrueI] show "PROP P" . |
|
21151 | 1351 |
next |
1352 |
assume "PROP P" |
|
23389 | 1353 |
then show "PROP P" . |
21151 | 1354 |
qed |
1355 |
||
59864 | 1356 |
lemma implies_True_equals: "(PROP P \<Longrightarrow> True) \<equiv> Trueprop True" |
61169 | 1357 |
by standard (intro TrueI) |
59864 | 1358 |
|
1359 |
lemma False_implies_equals: "(False \<Longrightarrow> P) \<equiv> Trueprop True" |
|
61169 | 1360 |
by standard simp_all |
59864 | 1361 |
|
71842
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1362 |
(* It seems that making this a simp rule is slower than using the simproc below *) |
63575 | 1363 |
lemma implies_False_swap: |
71842
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1364 |
"(False \<Longrightarrow> PROP P \<Longrightarrow> PROP Q) \<equiv> (PROP P \<Longrightarrow> False \<Longrightarrow> PROP Q)" |
63575 | 1365 |
by (rule swap_prems_eq) |
60169
5ef8ed685965
swap False to the right in assumptions to be eliminated at the right end
nipkow
parents:
60151
diff
changeset
|
1366 |
|
80663 | 1367 |
simproc_setup eliminate_false_implies ("False \<Longrightarrow> PROP P") = \<open> |
71842
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1368 |
let |
80663 | 1369 |
fun conv n = |
71842
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1370 |
if n > 1 then |
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1371 |
Conv.rewr_conv @{thm Pure.swap_prems_eq} |
80663 | 1372 |
then_conv Conv.arg_conv (conv (n - 1)) |
71842
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1373 |
then_conv Conv.rewr_conv @{thm HOL.implies_True_equals} |
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1374 |
else |
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1375 |
Conv.rewr_conv @{thm HOL.False_implies_equals} |
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1376 |
in |
80663 | 1377 |
fn _ => fn _ => fn ct => |
1378 |
let |
|
1379 |
val t = Thm.term_of ct |
|
1380 |
val n = length (Logic.strip_imp_prems t) |
|
1381 |
in |
|
1382 |
(case Logic.strip_imp_concl t of |
|
1383 |
\<^Const_>\<open>Trueprop for _\<close> => SOME (conv n ct) |
|
1384 |
| _ => NONE) |
|
1385 |
end |
|
71842
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1386 |
end |
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1387 |
\<close> |
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1388 |
|
db120661dded
new HOL simproc: eliminate_false_implies
Manuel Eberl <eberlm@in.tum.de>
parents:
71833
diff
changeset
|
1389 |
|
21151 | 1390 |
lemma ex_simps: |
60759 | 1391 |
"\<And>P Q. (\<exists>x. P x \<and> Q) = ((\<exists>x. P x) \<and> Q)" |
1392 |
"\<And>P Q. (\<exists>x. P \<and> Q x) = (P \<and> (\<exists>x. Q x))" |
|
1393 |
"\<And>P Q. (\<exists>x. P x \<or> Q) = ((\<exists>x. P x) \<or> Q)" |
|
1394 |
"\<And>P Q. (\<exists>x. P \<or> Q x) = (P \<or> (\<exists>x. Q x))" |
|
1395 |
"\<And>P Q. (\<exists>x. P x \<longrightarrow> Q) = ((\<forall>x. P x) \<longrightarrow> Q)" |
|
1396 |
"\<And>P Q. (\<exists>x. P \<longrightarrow> Q x) = (P \<longrightarrow> (\<exists>x. Q x))" |
|
61799 | 1397 |
\<comment> \<open>Miniscoping: pushing in existential quantifiers.\<close> |
21151 | 1398 |
by (iprover | blast)+ |
1399 |
||
1400 |
lemma all_simps: |
|
60759 | 1401 |
"\<And>P Q. (\<forall>x. P x \<and> Q) = ((\<forall>x. P x) \<and> Q)" |
1402 |
"\<And>P Q. (\<forall>x. P \<and> Q x) = (P \<and> (\<forall>x. Q x))" |
|
1403 |
"\<And>P Q. (\<forall>x. P x \<or> Q) = ((\<forall>x. P x) \<or> Q)" |
|
1404 |
"\<And>P Q. (\<forall>x. P \<or> Q x) = (P \<or> (\<forall>x. Q x))" |
|
1405 |
"\<And>P Q. (\<forall>x. P x \<longrightarrow> Q) = ((\<exists>x. P x) \<longrightarrow> Q)" |
|
1406 |
"\<And>P Q. (\<forall>x. P \<longrightarrow> Q x) = (P \<longrightarrow> (\<forall>x. Q x))" |
|
61799 | 1407 |
\<comment> \<open>Miniscoping: pushing in universal quantifiers.\<close> |
21151 | 1408 |
by (iprover | blast)+ |
15481 | 1409 |
|
21671 | 1410 |
lemmas [simp] = |
63575 | 1411 |
triv_forall_equality \<comment> \<open>prunes params\<close> |
1412 |
True_implies_equals implies_True_equals \<comment> \<open>prune \<open>True\<close> in asms\<close> |
|
1413 |
False_implies_equals \<comment> \<open>prune \<open>False\<close> in asms\<close> |
|
21671 | 1414 |
if_True |
1415 |
if_False |
|
1416 |
if_cancel |
|
1417 |
if_eq_cancel |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67405
diff
changeset
|
1418 |
imp_disjL \<comment> \<open>In general it seems wrong to add distributive laws by default: they |
63575 | 1419 |
might cause exponential blow-up. But \<open>imp_disjL\<close> has been in for a while |
20973 | 1420 |
and cannot be removed without affecting existing proofs. Moreover, |
63575 | 1421 |
rewriting by \<open>(P \<or> Q \<longrightarrow> R) = ((P \<longrightarrow> R) \<and> (Q \<longrightarrow> R))\<close> might be justified on the |
1422 |
grounds that it allows simplification of \<open>R\<close> in the two cases.\<close> |
|
21671 | 1423 |
conj_assoc |
1424 |
disj_assoc |
|
1425 |
de_Morgan_conj |
|
1426 |
de_Morgan_disj |
|
1427 |
imp_disj1 |
|
1428 |
imp_disj2 |
|
1429 |
not_imp |
|
1430 |
disj_not1 |
|
1431 |
not_all |
|
1432 |
not_ex |
|
1433 |
cases_simp |
|
1434 |
the_eq_trivial |
|
1435 |
the_sym_eq_trivial |
|
1436 |
ex_simps |
|
1437 |
all_simps |
|
1438 |
simp_thms |
|
71918 | 1439 |
subst_all |
21671 | 1440 |
|
1441 |
lemmas [cong] = imp_cong simp_implies_cong |
|
62390 | 1442 |
lemmas [split] = if_split |
20973 | 1443 |
|
69593 | 1444 |
ML \<open>val HOL_ss = simpset_of \<^context>\<close> |
20973 | 1445 |
|
63575 | 1446 |
text \<open>Simplifies \<open>x\<close> assuming \<open>c\<close> and \<open>y\<close> assuming \<open>\<not> c\<close>.\<close> |
20944 | 1447 |
lemma if_cong: |
1448 |
assumes "b = c" |
|
63575 | 1449 |
and "c \<Longrightarrow> x = u" |
1450 |
and "\<not> c \<Longrightarrow> y = v" |
|
20944 | 1451 |
shows "(if b then x else y) = (if c then u else v)" |
38525 | 1452 |
using assms by simp |
20944 | 1453 |
|
63575 | 1454 |
text \<open>Prevents simplification of \<open>x\<close> and \<open>y\<close>: |
60758 | 1455 |
faster and allows the execution of functional programs.\<close> |
20944 | 1456 |
lemma if_weak_cong [cong]: |
1457 |
assumes "b = c" |
|
1458 |
shows "(if b then x else y) = (if c then x else y)" |
|
23553 | 1459 |
using assms by (rule arg_cong) |
20944 | 1460 |
|
60758 | 1461 |
text \<open>Prevents simplification of t: much faster\<close> |
20944 | 1462 |
lemma let_weak_cong: |
1463 |
assumes "a = b" |
|
1464 |
shows "(let x = a in t x) = (let x = b in t x)" |
|
23553 | 1465 |
using assms by (rule arg_cong) |
20944 | 1466 |
|
60758 | 1467 |
text \<open>To tidy up the result of a simproc. Only the RHS will be simplified.\<close> |
20944 | 1468 |
lemma eq_cong2: |
1469 |
assumes "u = u'" |
|
1470 |
shows "(t \<equiv> u) \<equiv> (t \<equiv> u')" |
|
23553 | 1471 |
using assms by simp |
20944 | 1472 |
|
63575 | 1473 |
lemma if_distrib: "f (if c then x else y) = (if c then f x else f y)" |
20944 | 1474 |
by simp |
1475 |
||
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67719
diff
changeset
|
1476 |
lemma if_distribR: "(if b then f else g) x = (if b then f x else g x)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67719
diff
changeset
|
1477 |
by simp |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67719
diff
changeset
|
1478 |
|
67673
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1479 |
lemma all_if_distrib: "(\<forall>x. if x = a then P x else Q x) \<longleftrightarrow> P a \<and> (\<forall>x. x\<noteq>a \<longrightarrow> Q x)" |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1480 |
by auto |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1481 |
|
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1482 |
lemma ex_if_distrib: "(\<exists>x. if x = a then P x else Q x) \<longleftrightarrow> P a \<or> (\<exists>x. x\<noteq>a \<and> Q x)" |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1483 |
by auto |
c8caefb20564
lots of new material, ultimately related to measure theory
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
1484 |
|
67719
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
paulson <lp15@cam.ac.uk>
parents:
67673
diff
changeset
|
1485 |
lemma if_if_eq_conj: "(if P then if Q then x else y else y) = (if P \<and> Q then x else y)" |
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
paulson <lp15@cam.ac.uk>
parents:
67673
diff
changeset
|
1486 |
by simp |
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
paulson <lp15@cam.ac.uk>
parents:
67673
diff
changeset
|
1487 |
|
63575 | 1488 |
text \<open>As a simplification rule, it replaces all function equalities by |
60758 | 1489 |
first-order equalities.\<close> |
44277
bcb696533579
moved fundamental lemma fun_eq_iff to theory HOL; tuned whitespace
haftmann
parents:
44121
diff
changeset
|
1490 |
lemma fun_eq_iff: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)" |
bcb696533579
moved fundamental lemma fun_eq_iff to theory HOL; tuned whitespace
haftmann
parents:
44121
diff
changeset
|
1491 |
by auto |
bcb696533579
moved fundamental lemma fun_eq_iff to theory HOL; tuned whitespace
haftmann
parents:
44121
diff
changeset
|
1492 |
|
17459 | 1493 |
|
60758 | 1494 |
subsubsection \<open>Generic cases and induction\<close> |
17459 | 1495 |
|
60758 | 1496 |
text \<open>Rule projections:\<close> |
1497 |
ML \<open> |
|
32172 | 1498 |
structure Project_Rule = Project_Rule |
25388 | 1499 |
( |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1500 |
val conjunct1 = @{thm conjunct1} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1501 |
val conjunct2 = @{thm conjunct2} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1502 |
val mp = @{thm mp} |
59929 | 1503 |
); |
60758 | 1504 |
\<close> |
17459 | 1505 |
|
59940
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents:
59929
diff
changeset
|
1506 |
context |
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents:
59929
diff
changeset
|
1507 |
begin |
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents:
59929
diff
changeset
|
1508 |
|
59990
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59970
diff
changeset
|
1509 |
qualified definition "induct_forall P \<equiv> \<forall>x. P x" |
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59970
diff
changeset
|
1510 |
qualified definition "induct_implies A B \<equiv> A \<longrightarrow> B" |
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59970
diff
changeset
|
1511 |
qualified definition "induct_equal x y \<equiv> x = y" |
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59970
diff
changeset
|
1512 |
qualified definition "induct_conj A B \<equiv> A \<and> B" |
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59970
diff
changeset
|
1513 |
qualified definition "induct_true \<equiv> True" |
a81dc82ecba3
clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
wenzelm
parents:
59970
diff
changeset
|
1514 |
qualified definition "induct_false \<equiv> False" |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35115
diff
changeset
|
1515 |
|
59929 | 1516 |
lemma induct_forall_eq: "(\<And>x. P x) \<equiv> Trueprop (induct_forall (\<lambda>x. P x))" |
18457 | 1517 |
by (unfold atomize_all induct_forall_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1518 |
|
59929 | 1519 |
lemma induct_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop (induct_implies A B)" |
18457 | 1520 |
by (unfold atomize_imp induct_implies_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1521 |
|
59929 | 1522 |
lemma induct_equal_eq: "(x \<equiv> y) \<equiv> Trueprop (induct_equal x y)" |
18457 | 1523 |
by (unfold atomize_eq induct_equal_def) |
1524 |
||
59929 | 1525 |
lemma induct_conj_eq: "(A &&& B) \<equiv> Trueprop (induct_conj A B)" |
18457 | 1526 |
by (unfold atomize_conj induct_conj_def) |
1527 |
||
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1528 |
lemmas induct_atomize' = induct_forall_eq induct_implies_eq induct_conj_eq |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1529 |
lemmas induct_atomize = induct_atomize' induct_equal_eq |
45607 | 1530 |
lemmas induct_rulify' [symmetric] = induct_atomize' |
1531 |
lemmas induct_rulify [symmetric] = induct_atomize |
|
18457 | 1532 |
lemmas induct_rulify_fallback = |
1533 |
induct_forall_def induct_implies_def induct_equal_def induct_conj_def |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1534 |
induct_true_def induct_false_def |
18457 | 1535 |
|
11989 | 1536 |
lemma induct_forall_conj: "induct_forall (\<lambda>x. induct_conj (A x) (B x)) = |
1537 |
induct_conj (induct_forall A) (induct_forall B)" |
|
17589 | 1538 |
by (unfold induct_forall_def induct_conj_def) iprover |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1539 |
|
11989 | 1540 |
lemma induct_implies_conj: "induct_implies C (induct_conj A B) = |
1541 |
induct_conj (induct_implies C A) (induct_implies C B)" |
|
17589 | 1542 |
by (unfold induct_implies_def induct_conj_def) iprover |
11989 | 1543 |
|
59929 | 1544 |
lemma induct_conj_curry: "(induct_conj A B \<Longrightarrow> PROP C) \<equiv> (A \<Longrightarrow> B \<Longrightarrow> PROP C)" |
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1545 |
proof |
59929 | 1546 |
assume r: "induct_conj A B \<Longrightarrow> PROP C" |
1547 |
assume ab: A B |
|
1548 |
show "PROP C" by (rule r) (simp add: induct_conj_def ab) |
|
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1549 |
next |
59929 | 1550 |
assume r: "A \<Longrightarrow> B \<Longrightarrow> PROP C" |
1551 |
assume ab: "induct_conj A B" |
|
1552 |
show "PROP C" by (rule r) (simp_all add: ab [unfolded induct_conj_def]) |
|
13598
8bc77b17f59f
Fixed problem with induct_conj_curry: variable C should have type prop.
berghofe
parents:
13596
diff
changeset
|
1553 |
qed |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1554 |
|
11989 | 1555 |
lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1556 |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1557 |
lemma induct_trueI: "induct_true" |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1558 |
by (simp add: induct_true_def) |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1559 |
|
60758 | 1560 |
text \<open>Method setup.\<close> |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1561 |
|
69605 | 1562 |
ML_file \<open>~~/src/Tools/induct.ML\<close> |
60758 | 1563 |
ML \<open> |
32171 | 1564 |
structure Induct = Induct |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1565 |
( |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1566 |
val cases_default = @{thm case_split} |
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1567 |
val atomize = @{thms induct_atomize} |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1568 |
val rulify = @{thms induct_rulify'} |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1569 |
val rulify_fallback = @{thms induct_rulify_fallback} |
34988
cca208c8d619
Added setup for simplification of equality constraints in cases rules.
berghofe
parents:
34917
diff
changeset
|
1570 |
val equal_def = @{thm induct_equal_def} |
80662 | 1571 |
fun dest_def \<^Const_>\<open>induct_equal _ for t u\<close> = SOME (t, u) |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1572 |
| dest_def _ = NONE |
58957 | 1573 |
fun trivial_tac ctxt = match_tac ctxt @{thms induct_trueI} |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1574 |
) |
60758 | 1575 |
\<close> |
11824
f4c1882dde2c
setup generic cases and induction (from Inductive.thy);
wenzelm
parents:
11770
diff
changeset
|
1576 |
|
69605 | 1577 |
ML_file \<open>~~/src/Tools/induction.ML\<close> |
45014
0e847655b2d8
New proof method "induction" that gives induction hypotheses the name IH.
nipkow
parents:
44921
diff
changeset
|
1578 |
|
78799 | 1579 |
simproc_setup passive swap_induct_false ("induct_false \<Longrightarrow> PROP P \<Longrightarrow> PROP Q") = |
78794 | 1580 |
\<open>fn _ => fn _ => fn ct => |
1581 |
(case Thm.term_of ct of |
|
1582 |
_ $ (P as _ $ \<^Const_>\<open>induct_false\<close>) $ (_ $ Q $ _) => |
|
1583 |
if P <> Q then SOME Drule.swap_prems_eq else NONE |
|
1584 |
| _ => NONE)\<close> |
|
78799 | 1585 |
|
1586 |
simproc_setup passive induct_equal_conj_curry ("induct_conj P Q \<Longrightarrow> PROP R") = |
|
78794 | 1587 |
\<open>fn _ => fn _ => fn ct => |
1588 |
(case Thm.term_of ct of |
|
1589 |
_ $ (_ $ P) $ _ => |
|
1590 |
let |
|
1591 |
fun is_conj \<^Const_>\<open>induct_conj for P Q\<close> = is_conj P andalso is_conj Q |
|
1592 |
| is_conj \<^Const_>\<open>induct_equal _ for _ _\<close> = true |
|
1593 |
| is_conj \<^Const_>\<open>induct_true\<close> = true |
|
1594 |
| is_conj \<^Const_>\<open>induct_false\<close> = true |
|
1595 |
| is_conj _ = false |
|
1596 |
in if is_conj P then SOME @{thm induct_conj_curry} else NONE end |
|
1597 |
| _ => NONE)\<close> |
|
1598 |
||
60758 | 1599 |
declaration \<open> |
80701 | 1600 |
K (Induct.map_simpset |
1601 |
(Simplifier.add_proc \<^simproc>\<open>swap_induct_false\<close> #> |
|
1602 |
Simplifier.add_proc \<^simproc>\<open>induct_equal_conj_curry\<close> #> |
|
1603 |
Simplifier.set_mksimps (fn ctxt => |
|
1604 |
Simpdata.mksimps Simpdata.mksimps_pairs ctxt #> |
|
1605 |
map (rewrite_rule ctxt (map Thm.symmetric @{thms induct_rulify_fallback}))))) |
|
60758 | 1606 |
\<close> |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1607 |
|
60758 | 1608 |
text \<open>Pre-simplification of induction and cases rules\<close> |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1609 |
|
59929 | 1610 |
lemma [induct_simp]: "(\<And>x. induct_equal x t \<Longrightarrow> PROP P x) \<equiv> PROP P t" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1611 |
unfolding induct_equal_def |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1612 |
proof |
59929 | 1613 |
assume r: "\<And>x. x = t \<Longrightarrow> PROP P x" |
1614 |
show "PROP P t" by (rule r [OF refl]) |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1615 |
next |
59929 | 1616 |
fix x |
1617 |
assume "PROP P t" "x = t" |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1618 |
then show "PROP P x" by simp |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1619 |
qed |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1620 |
|
59929 | 1621 |
lemma [induct_simp]: "(\<And>x. induct_equal t x \<Longrightarrow> PROP P x) \<equiv> PROP P t" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1622 |
unfolding induct_equal_def |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1623 |
proof |
59929 | 1624 |
assume r: "\<And>x. t = x \<Longrightarrow> PROP P x" |
1625 |
show "PROP P t" by (rule r [OF refl]) |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1626 |
next |
59929 | 1627 |
fix x |
1628 |
assume "PROP P t" "t = x" |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1629 |
then show "PROP P x" by simp |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1630 |
qed |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1631 |
|
59929 | 1632 |
lemma [induct_simp]: "(induct_false \<Longrightarrow> P) \<equiv> Trueprop induct_true" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1633 |
unfolding induct_false_def induct_true_def |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1634 |
by (iprover intro: equal_intr_rule) |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1635 |
|
59929 | 1636 |
lemma [induct_simp]: "(induct_true \<Longrightarrow> PROP P) \<equiv> PROP P" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1637 |
unfolding induct_true_def |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1638 |
proof |
59929 | 1639 |
assume "True \<Longrightarrow> PROP P" |
1640 |
then show "PROP P" using TrueI . |
|
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1641 |
next |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1642 |
assume "PROP P" |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1643 |
then show "PROP P" . |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1644 |
qed |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1645 |
|
59929 | 1646 |
lemma [induct_simp]: "(PROP P \<Longrightarrow> induct_true) \<equiv> Trueprop induct_true" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1647 |
unfolding induct_true_def |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1648 |
by (iprover intro: equal_intr_rule) |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1649 |
|
62958
b41c1cb5e251
Type_Infer.object_logic controls improvement of type inference result;
wenzelm
parents:
62913
diff
changeset
|
1650 |
lemma [induct_simp]: "(\<And>x::'a::{}. induct_true) \<equiv> Trueprop induct_true" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1651 |
unfolding induct_true_def |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1652 |
by (iprover intro: equal_intr_rule) |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1653 |
|
59929 | 1654 |
lemma [induct_simp]: "induct_implies induct_true P \<equiv> P" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1655 |
by (simp add: induct_implies_def induct_true_def) |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1656 |
|
59929 | 1657 |
lemma [induct_simp]: "x = x \<longleftrightarrow> True" |
34908
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1658 |
by (rule simp_thms) |
d546e75631bb
Added setup for simplification of equality constraints in induction rules.
berghofe
parents:
34294
diff
changeset
|
1659 |
|
59940
087d81f5213e
local setup of induction tools, with restricted access to auxiliary consts;
wenzelm
parents:
59929
diff
changeset
|
1660 |
end |
18457 | 1661 |
|
69605 | 1662 |
ML_file \<open>~~/src/Tools/induct_tacs.ML\<close> |
27126
3ede9103de8e
eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents:
27107
diff
changeset
|
1663 |
|
20944 | 1664 |
|
60758 | 1665 |
subsubsection \<open>Coherent logic\<close> |
28325 | 1666 |
|
69605 | 1667 |
ML_file \<open>~~/src/Tools/coherent.ML\<close> |
60758 | 1668 |
ML \<open> |
32734 | 1669 |
structure Coherent = Coherent |
28325 | 1670 |
( |
55632 | 1671 |
val atomize_elimL = @{thm atomize_elimL}; |
1672 |
val atomize_exL = @{thm atomize_exL}; |
|
1673 |
val atomize_conjL = @{thm atomize_conjL}; |
|
1674 |
val atomize_disjL = @{thm atomize_disjL}; |
|
69593 | 1675 |
val operator_names = [\<^const_name>\<open>HOL.disj\<close>, \<^const_name>\<open>HOL.conj\<close>, \<^const_name>\<open>Ex\<close>]; |
28325 | 1676 |
); |
60758 | 1677 |
\<close> |
28325 | 1678 |
|
1679 |
||
60758 | 1680 |
subsubsection \<open>Reorienting equalities\<close> |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1681 |
|
60758 | 1682 |
ML \<open> |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1683 |
signature REORIENT_PROC = |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1684 |
sig |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1685 |
val add : (term -> bool) -> theory -> theory |
78800 | 1686 |
val proc : Simplifier.proc |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1687 |
end; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1688 |
|
33523 | 1689 |
structure Reorient_Proc : REORIENT_PROC = |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1690 |
struct |
33523 | 1691 |
structure Data = Theory_Data |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1692 |
( |
33523 | 1693 |
type T = ((term -> bool) * stamp) list; |
1694 |
val empty = []; |
|
67405
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents:
67399
diff
changeset
|
1695 |
fun merge data : T = Library.merge (eq_snd (op =)) data; |
33523 | 1696 |
); |
1697 |
fun add m = Data.map (cons (m, stamp ())); |
|
1698 |
fun matches thy t = exists (fn (m, _) => m t) (Data.get thy); |
|
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1699 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1700 |
val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; |
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1701 |
fun proc ctxt ct = |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1702 |
let |
42361 | 1703 |
val thy = Proof_Context.theory_of ctxt; |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1704 |
in |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1705 |
case Thm.term_of ct of |
33523 | 1706 |
(_ $ t $ u) => if matches thy u then NONE else SOME meta_reorient |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1707 |
| _ => NONE |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1708 |
end; |
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1709 |
end; |
60758 | 1710 |
\<close> |
31024
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1711 |
|
0fdf666e08bf
reimplement reorientation simproc using theory data
huffman
parents:
30980
diff
changeset
|
1712 |
|
60758 | 1713 |
subsection \<open>Other simple lemmas and lemma duplicates\<close> |
20944 | 1714 |
|
74741
6e1fad4f602b
added eq_iff_swap for creating symmetric variants of thms; applied it in List.
nipkow
parents:
74561
diff
changeset
|
1715 |
lemma eq_iff_swap: "(x = y \<longleftrightarrow> P) \<Longrightarrow> (y = x \<longleftrightarrow> P)" |
6e1fad4f602b
added eq_iff_swap for creating symmetric variants of thms; applied it in List.
nipkow
parents:
74561
diff
changeset
|
1716 |
by blast |
6e1fad4f602b
added eq_iff_swap for creating symmetric variants of thms; applied it in List.
nipkow
parents:
74561
diff
changeset
|
1717 |
|
68975
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68072
diff
changeset
|
1718 |
lemma all_cong1: "(\<And>x. P x = P' x) \<Longrightarrow> (\<forall>x. P x) = (\<forall>x. P' x)" |
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68072
diff
changeset
|
1719 |
by auto |
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68072
diff
changeset
|
1720 |
|
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68072
diff
changeset
|
1721 |
lemma ex_cong1: "(\<And>x. P x = P' x) \<Longrightarrow> (\<exists>x. P x) = (\<exists>x. P' x)" |
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68072
diff
changeset
|
1722 |
by auto |
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68072
diff
changeset
|
1723 |
|
67091 | 1724 |
lemma all_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> (\<forall>x. Q x \<longrightarrow> P x) = (\<forall>x. Q x \<longrightarrow> P' x)" |
66836 | 1725 |
by auto |
1726 |
||
67091 | 1727 |
lemma ex_cong: "(\<And>x. Q x \<Longrightarrow> P x = P' x) \<Longrightarrow> (\<exists>x. Q x \<and> P x) = (\<exists>x. Q x \<and> P' x)" |
66836 | 1728 |
by auto |
1729 |
||
60759 | 1730 |
lemma ex1_eq [iff]: "\<exists>!x. x = t" "\<exists>!x. t = x" |
20944 | 1731 |
by blast+ |
1732 |
||
71608
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1733 |
lemma choice_eq: "(\<forall>x. \<exists>!y. P x y) = (\<exists>!f. \<forall>x. P x (f x))" (is "?lhs = ?rhs") |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1734 |
proof (intro iffI allI) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1735 |
assume L: ?lhs |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1736 |
then have \<section>: "\<forall>x. P x (THE y. P x y)" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1737 |
by (best intro: theI') |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1738 |
show ?rhs |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1739 |
by (rule ex1I) (use L \<section> in \<open>fast+\<close>) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1740 |
next |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1741 |
fix x |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1742 |
assume R: ?rhs |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1743 |
then obtain f where f: "\<forall>x. P x (f x)" and f1: "\<And>y. (\<forall>x. P x (y x)) \<Longrightarrow> y = f" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1744 |
by (blast elim: ex1E) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1745 |
show "\<exists>!y. P x y" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1746 |
proof (rule ex1I) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1747 |
show "P x (f x)" |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1748 |
using f by blast |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1749 |
show "y = f x" if "P x y" for y |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1750 |
proof - |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1751 |
have "P z (if z = x then y else f z)" for z |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1752 |
using f that by (auto split: if_split) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1753 |
with f1 [of "\<lambda>z. if z = x then y else f z"] f |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1754 |
show ?thesis |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1755 |
by (auto simp add: split: if_split_asm dest: fun_cong) |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1756 |
qed |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1757 |
qed |
856c68ab6f13
structured a lot of ancient, horrible proofs
paulson <lp15@cam.ac.uk>
parents:
71517
diff
changeset
|
1758 |
qed |
20944 | 1759 |
|
22218 | 1760 |
lemmas eq_sym_conv = eq_commute |
1761 |
||
23037
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1762 |
lemma nnf_simps: |
63575 | 1763 |
"(\<not> (P \<and> Q)) = (\<not> P \<or> \<not> Q)" |
1764 |
"(\<not> (P \<or> Q)) = (\<not> P \<and> \<not> Q)" |
|
1765 |
"(P \<longrightarrow> Q) = (\<not> P \<or> Q)" |
|
1766 |
"(P = Q) = ((P \<and> Q) \<or> (\<not> P \<and> \<not> Q))" |
|
1767 |
"(\<not> (P = Q)) = ((P \<and> \<not> Q) \<or> (\<not> P \<and> Q))" |
|
1768 |
"(\<not> \<not> P) = P" |
|
1769 |
by blast+ |
|
1770 |
||
23037
6c72943a71b1
added a set of NNF normalization lemmas and nnf_conv
chaieb
parents:
22993
diff
changeset
|
1771 |
|
60758 | 1772 |
subsection \<open>Basic ML bindings\<close> |
21671 | 1773 |
|
60758 | 1774 |
ML \<open> |
22129 | 1775 |
val FalseE = @{thm FalseE} |
1776 |
val Let_def = @{thm Let_def} |
|
1777 |
val TrueI = @{thm TrueI} |
|
1778 |
val allE = @{thm allE} |
|
1779 |
val allI = @{thm allI} |
|
1780 |
val all_dupE = @{thm all_dupE} |
|
1781 |
val arg_cong = @{thm arg_cong} |
|
1782 |
val box_equals = @{thm box_equals} |
|
1783 |
val ccontr = @{thm ccontr} |
|
1784 |
val classical = @{thm classical} |
|
1785 |
val conjE = @{thm conjE} |
|
1786 |
val conjI = @{thm conjI} |
|
1787 |
val conjunct1 = @{thm conjunct1} |
|
1788 |
val conjunct2 = @{thm conjunct2} |
|
1789 |
val disjCI = @{thm disjCI} |
|
1790 |
val disjE = @{thm disjE} |
|
1791 |
val disjI1 = @{thm disjI1} |
|
1792 |
val disjI2 = @{thm disjI2} |
|
1793 |
val eq_reflection = @{thm eq_reflection} |
|
1794 |
val ex1E = @{thm ex1E} |
|
1795 |
val ex1I = @{thm ex1I} |
|
1796 |
val ex1_implies_ex = @{thm ex1_implies_ex} |
|
1797 |
val exE = @{thm exE} |
|
1798 |
val exI = @{thm exI} |
|
1799 |
val excluded_middle = @{thm excluded_middle} |
|
1800 |
val ext = @{thm ext} |
|
1801 |
val fun_cong = @{thm fun_cong} |
|
1802 |
val iffD1 = @{thm iffD1} |
|
1803 |
val iffD2 = @{thm iffD2} |
|
1804 |
val iffI = @{thm iffI} |
|
1805 |
val impE = @{thm impE} |
|
1806 |
val impI = @{thm impI} |
|
1807 |
val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq} |
|
1808 |
val mp = @{thm mp} |
|
1809 |
val notE = @{thm notE} |
|
1810 |
val notI = @{thm notI} |
|
1811 |
val not_all = @{thm not_all} |
|
1812 |
val not_ex = @{thm not_ex} |
|
1813 |
val not_iff = @{thm not_iff} |
|
1814 |
val not_not = @{thm not_not} |
|
1815 |
val not_sym = @{thm not_sym} |
|
1816 |
val refl = @{thm refl} |
|
1817 |
val rev_mp = @{thm rev_mp} |
|
1818 |
val spec = @{thm spec} |
|
1819 |
val ssubst = @{thm ssubst} |
|
1820 |
val subst = @{thm subst} |
|
1821 |
val sym = @{thm sym} |
|
1822 |
val trans = @{thm trans} |
|
60758 | 1823 |
\<close> |
21671 | 1824 |
|
70486 | 1825 |
locale cnf |
1826 |
begin |
|
1827 |
||
1828 |
lemma clause2raw_notE: "\<lbrakk>P; \<not>P\<rbrakk> \<Longrightarrow> False" by auto |
|
1829 |
lemma clause2raw_not_disj: "\<lbrakk>\<not> P; \<not> Q\<rbrakk> \<Longrightarrow> \<not> (P \<or> Q)" by auto |
|
1830 |
lemma clause2raw_not_not: "P \<Longrightarrow> \<not>\<not> P" by auto |
|
1831 |
||
1832 |
lemma iff_refl: "(P::bool) = P" by auto |
|
1833 |
lemma iff_trans: "[| (P::bool) = Q; Q = R |] ==> P = R" by auto |
|
1834 |
lemma conj_cong: "[| P = P'; Q = Q' |] ==> (P \<and> Q) = (P' \<and> Q')" by auto |
|
1835 |
lemma disj_cong: "[| P = P'; Q = Q' |] ==> (P \<or> Q) = (P' \<or> Q')" by auto |
|
1836 |
||
1837 |
lemma make_nnf_imp: "[| (\<not>P) = P'; Q = Q' |] ==> (P \<longrightarrow> Q) = (P' \<or> Q')" by auto |
|
1838 |
lemma make_nnf_iff: "[| P = P'; (\<not>P) = NP; Q = Q'; (\<not>Q) = NQ |] ==> (P = Q) = ((P' \<or> NQ) \<and> (NP \<or> Q'))" by auto |
|
1839 |
lemma make_nnf_not_false: "(\<not>False) = True" by auto |
|
1840 |
lemma make_nnf_not_true: "(\<not>True) = False" by auto |
|
1841 |
lemma make_nnf_not_conj: "[| (\<not>P) = P'; (\<not>Q) = Q' |] ==> (\<not>(P \<and> Q)) = (P' \<or> Q')" by auto |
|
1842 |
lemma make_nnf_not_disj: "[| (\<not>P) = P'; (\<not>Q) = Q' |] ==> (\<not>(P \<or> Q)) = (P' \<and> Q')" by auto |
|
1843 |
lemma make_nnf_not_imp: "[| P = P'; (\<not>Q) = Q' |] ==> (\<not>(P \<longrightarrow> Q)) = (P' \<and> Q')" by auto |
|
1844 |
lemma make_nnf_not_iff: "[| P = P'; (\<not>P) = NP; Q = Q'; (\<not>Q) = NQ |] ==> (\<not>(P = Q)) = ((P' \<or> Q') \<and> (NP \<or> NQ))" by auto |
|
1845 |
lemma make_nnf_not_not: "P = P' ==> (\<not>\<not>P) = P'" by auto |
|
1846 |
||
1847 |
lemma simp_TF_conj_True_l: "[| P = True; Q = Q' |] ==> (P \<and> Q) = Q'" by auto |
|
1848 |
lemma simp_TF_conj_True_r: "[| P = P'; Q = True |] ==> (P \<and> Q) = P'" by auto |
|
1849 |
lemma simp_TF_conj_False_l: "P = False ==> (P \<and> Q) = False" by auto |
|
1850 |
lemma simp_TF_conj_False_r: "Q = False ==> (P \<and> Q) = False" by auto |
|
1851 |
lemma simp_TF_disj_True_l: "P = True ==> (P \<or> Q) = True" by auto |
|
1852 |
lemma simp_TF_disj_True_r: "Q = True ==> (P \<or> Q) = True" by auto |
|
1853 |
lemma simp_TF_disj_False_l: "[| P = False; Q = Q' |] ==> (P \<or> Q) = Q'" by auto |
|
1854 |
lemma simp_TF_disj_False_r: "[| P = P'; Q = False |] ==> (P \<or> Q) = P'" by auto |
|
1855 |
||
1856 |
lemma make_cnf_disj_conj_l: "[| (P \<or> R) = PR; (Q \<or> R) = QR |] ==> ((P \<and> Q) \<or> R) = (PR \<and> QR)" by auto |
|
1857 |
lemma make_cnf_disj_conj_r: "[| (P \<or> Q) = PQ; (P \<or> R) = PR |] ==> (P \<or> (Q \<and> R)) = (PQ \<and> PR)" by auto |
|
1858 |
||
1859 |
lemma make_cnfx_disj_ex_l: "((\<exists>(x::bool). P x) \<or> Q) = (\<exists>x. P x \<or> Q)" by auto |
|
1860 |
lemma make_cnfx_disj_ex_r: "(P \<or> (\<exists>(x::bool). Q x)) = (\<exists>x. P \<or> Q x)" by auto |
|
1861 |
lemma make_cnfx_newlit: "(P \<or> Q) = (\<exists>x. (P \<or> x) \<and> (Q \<or> \<not>x))" by auto |
|
1862 |
lemma make_cnfx_ex_cong: "(\<forall>(x::bool). P x = Q x) \<Longrightarrow> (\<exists>x. P x) = (\<exists>x. Q x)" by auto |
|
1863 |
||
1864 |
lemma weakening_thm: "[| P; Q |] ==> Q" by auto |
|
1865 |
||
1866 |
lemma cnftac_eq_imp: "[| P = Q; P |] ==> Q" by auto |
|
1867 |
||
1868 |
end |
|
1869 |
||
69605 | 1870 |
ML_file \<open>Tools/cnf.ML\<close> |
55239 | 1871 |
|
21671 | 1872 |
|
61799 | 1873 |
section \<open>\<open>NO_MATCH\<close> simproc\<close> |
58775
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1874 |
|
60758 | 1875 |
text \<open> |
63575 | 1876 |
The simplification procedure can be used to avoid simplification of terms |
1877 |
of a certain form. |
|
60758 | 1878 |
\<close> |
58775
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1879 |
|
63575 | 1880 |
definition NO_MATCH :: "'a \<Rightarrow> 'b \<Rightarrow> bool" |
1881 |
where "NO_MATCH pat val \<equiv> True" |
|
58830 | 1882 |
|
63575 | 1883 |
lemma NO_MATCH_cong[cong]: "NO_MATCH pat val = NO_MATCH pat val" |
1884 |
by (rule refl) |
|
58775
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1885 |
|
58830 | 1886 |
declare [[coercion_args NO_MATCH - -]] |
1887 |
||
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1888 |
simproc_setup NO_MATCH ("NO_MATCH pat val") = \<open>K (fn ctxt => fn ct => |
58775
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1889 |
let |
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1890 |
val thy = Proof_Context.theory_of ctxt |
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1891 |
val dest_binop = Term.dest_comb #> apfst (Term.dest_comb #> snd) |
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1892 |
val m = Pattern.matches thy (dest_binop (Thm.term_of ct)) |
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
77172
diff
changeset
|
1893 |
in if m then NONE else SOME @{thm NO_MATCH_def} end) |
60758 | 1894 |
\<close> |
58775
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1895 |
|
60758 | 1896 |
text \<open> |
69593 | 1897 |
This setup ensures that a rewrite rule of the form \<^term>\<open>NO_MATCH pat val \<Longrightarrow> t\<close> |
63575 | 1898 |
is only applied, if the pattern \<open>pat\<close> does not match the value \<open>val\<close>. |
60758 | 1899 |
\<close> |
58775
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1900 |
|
9cd64a66a765
move NO_MATCH simproc from the AFP entry Graph_Theory to HOL
hoelzl
parents:
58659
diff
changeset
|
1901 |
|
63575 | 1902 |
text\<open> |
1903 |
Tagging a premise of a simp rule with ASSUMPTION forces the simplifier |
|
1904 |
not to simplify the argument and to solve it by an assumption. |
|
1905 |
\<close> |
|
61202 | 1906 |
|
63575 | 1907 |
definition ASSUMPTION :: "bool \<Rightarrow> bool" |
1908 |
where "ASSUMPTION A \<equiv> A" |
|
61202 | 1909 |
|
1910 |
lemma ASSUMPTION_cong[cong]: "ASSUMPTION A = ASSUMPTION A" |
|
63575 | 1911 |
by (rule refl) |
61202 | 1912 |
|
1913 |
lemma ASSUMPTION_I: "A \<Longrightarrow> ASSUMPTION A" |
|
63575 | 1914 |
by (simp add: ASSUMPTION_def) |
61202 | 1915 |
|
1916 |
lemma ASSUMPTION_D: "ASSUMPTION A \<Longrightarrow> A" |
|
63575 | 1917 |
by (simp add: ASSUMPTION_def) |
61202 | 1918 |
|
61222 | 1919 |
setup \<open> |
82661 | 1920 |
(map_theory_simpset o Simplifier.add_unsafe_solver) ( |
1921 |
mk_solver "ASSUMPTION" (fn ctxt => |
|
1922 |
resolve_tac ctxt @{thms ASSUMPTION_I} THEN' |
|
1923 |
resolve_tac ctxt (Simplifier.prems_of ctxt))) |
|
61222 | 1924 |
\<close> |
61202 | 1925 |
|
1926 |
||
60758 | 1927 |
subsection \<open>Code generator setup\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1928 |
|
60758 | 1929 |
subsubsection \<open>Generic code generator preprocessor setup\<close> |
31151 | 1930 |
|
63575 | 1931 |
lemma conj_left_cong: "P \<longleftrightarrow> Q \<Longrightarrow> P \<and> R \<longleftrightarrow> Q \<and> R" |
53146
3a93bc5d3370
congruence rules for code_simp to mimic typical non-strict behaviour of conj and disj
haftmann
parents:
52654
diff
changeset
|
1932 |
by (fact arg_cong) |
3a93bc5d3370
congruence rules for code_simp to mimic typical non-strict behaviour of conj and disj
haftmann
parents:
52654
diff
changeset
|
1933 |
|
63575 | 1934 |
lemma disj_left_cong: "P \<longleftrightarrow> Q \<Longrightarrow> P \<or> R \<longleftrightarrow> Q \<or> R" |
53146
3a93bc5d3370
congruence rules for code_simp to mimic typical non-strict behaviour of conj and disj
haftmann
parents:
52654
diff
changeset
|
1935 |
by (fact arg_cong) |
3a93bc5d3370
congruence rules for code_simp to mimic typical non-strict behaviour of conj and disj
haftmann
parents:
52654
diff
changeset
|
1936 |
|
60758 | 1937 |
setup \<open> |
58826 | 1938 |
Code_Preproc.map_pre (put_simpset HOL_basic_ss) #> |
1939 |
Code_Preproc.map_post (put_simpset HOL_basic_ss) #> |
|
1940 |
Code_Simp.map_ss (put_simpset HOL_basic_ss #> |
|
1941 |
Simplifier.add_cong @{thm conj_left_cong} #> |
|
1942 |
Simplifier.add_cong @{thm disj_left_cong}) |
|
60758 | 1943 |
\<close> |
31151 | 1944 |
|
53146
3a93bc5d3370
congruence rules for code_simp to mimic typical non-strict behaviour of conj and disj
haftmann
parents:
52654
diff
changeset
|
1945 |
|
60758 | 1946 |
subsubsection \<open>Equality\<close> |
24844
98c006a30218
certificates for code generator case expressions
haftmann
parents:
24842
diff
changeset
|
1947 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1948 |
class equal = |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1949 |
fixes equal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1950 |
assumes equal_eq: "equal x y \<longleftrightarrow> x = y" |
26513 | 1951 |
begin |
1952 |
||
67399 | 1953 |
lemma equal: "equal = (=)" |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1954 |
by (rule ext equal_eq)+ |
28346
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1955 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1956 |
lemma equal_refl: "equal x x \<longleftrightarrow> True" |
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74741
diff
changeset
|
1957 |
unfolding equal by (rule iffI TrueI refl)+ |
28346
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
haftmann
parents:
28325
diff
changeset
|
1958 |
|
67399 | 1959 |
lemma eq_equal: "(=) \<equiv> equal" |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1960 |
by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq) |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1961 |
|
26513 | 1962 |
end |
1963 |
||
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1964 |
declare eq_equal [symmetric, code_post] |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
1965 |
declare eq_equal [code] |
30966 | 1966 |
|
78799 | 1967 |
simproc_setup passive equal (HOL.eq) = |
78794 | 1968 |
\<open>fn _ => fn _ => fn ct => |
1969 |
(case Thm.term_of ct of |
|
80662 | 1970 |
\<^Const_>\<open>HOL.eq T\<close> => if is_Type T then SOME @{thm eq_equal} else NONE |
78794 | 1971 |
| _ => NONE)\<close> |
1972 |
||
80701 | 1973 |
setup \<open>Code_Preproc.map_pre (Simplifier.add_proc \<^simproc>\<open>equal\<close>)\<close> |
31151 | 1974 |
|
30966 | 1975 |
|
60758 | 1976 |
subsubsection \<open>Generic code generator foundation\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1977 |
|
69593 | 1978 |
text \<open>Datatype \<^typ>\<open>bool\<close>\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1979 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1980 |
code_datatype True False |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1981 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1982 |
lemma [code]: |
33185
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1983 |
shows "False \<and> P \<longleftrightarrow> False" |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1984 |
and "True \<and> P \<longleftrightarrow> P" |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1985 |
and "P \<and> False \<longleftrightarrow> False" |
63575 | 1986 |
and "P \<and> True \<longleftrightarrow> P" |
1987 |
by simp_all |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1988 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1989 |
lemma [code]: |
33185
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1990 |
shows "False \<or> P \<longleftrightarrow> P" |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1991 |
and "True \<or> P \<longleftrightarrow> True" |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1992 |
and "P \<or> False \<longleftrightarrow> P" |
63575 | 1993 |
and "P \<or> True \<longleftrightarrow> True" |
1994 |
by simp_all |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
1995 |
|
33185
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1996 |
lemma [code]: |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1997 |
shows "(False \<longrightarrow> P) \<longleftrightarrow> True" |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1998 |
and "(True \<longrightarrow> P) \<longleftrightarrow> P" |
247f6c6969d9
tuned code setup for primitive boolean connectors
haftmann
parents:
33084
diff
changeset
|
1999 |
and "(P \<longrightarrow> False) \<longleftrightarrow> \<not> P" |
63575 | 2000 |
and "(P \<longrightarrow> True) \<longleftrightarrow> True" |
2001 |
by simp_all |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2002 |
|
69593 | 2003 |
text \<open>More about \<^typ>\<open>prop\<close>\<close> |
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2004 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2005 |
lemma [code nbe]: |
58826 | 2006 |
shows "(True \<Longrightarrow> PROP Q) \<equiv> PROP Q" |
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2007 |
and "(PROP Q \<Longrightarrow> True) \<equiv> Trueprop True" |
63575 | 2008 |
and "(P \<Longrightarrow> R) \<equiv> Trueprop (P \<longrightarrow> R)" |
2009 |
by (auto intro!: equal_intr_rule) |
|
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2010 |
|
63575 | 2011 |
lemma Trueprop_code [code]: "Trueprop True \<equiv> Code_Generator.holds" |
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2012 |
by (auto intro!: equal_intr_rule holds) |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2013 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2014 |
declare Trueprop_code [symmetric, code_post] |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2015 |
|
60758 | 2016 |
text \<open>Equality\<close> |
39421
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2017 |
|
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2018 |
declare simp_thms(6) [code nbe] |
b6a77cffc231
introduced "holds" as synthetic datatype constructor for "prop"; moved Pure code generator setup to Code_Generator.thy
haftmann
parents:
39403
diff
changeset
|
2019 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
2020 |
instantiation itself :: (type) equal |
31132 | 2021 |
begin |
2022 |
||
63575 | 2023 |
definition equal_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" |
2024 |
where "equal_itself x y \<longleftrightarrow> x = y" |
|
31132 | 2025 |
|
63575 | 2026 |
instance |
2027 |
by standard (fact equal_itself_def) |
|
31132 | 2028 |
|
2029 |
end |
|
2030 |
||
63575 | 2031 |
lemma equal_itself_code [code]: "equal TYPE('a) TYPE('a) \<longleftrightarrow> True" |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
2032 |
by (simp add: equal) |
31132 | 2033 |
|
69593 | 2034 |
setup \<open>Sign.add_const_constraint (\<^const_name>\<open>equal\<close>, SOME \<^typ>\<open>'a::type \<Rightarrow> 'a \<Rightarrow> bool\<close>)\<close> |
31956
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2035 |
|
67399 | 2036 |
lemma equal_alias_cert: "OFCLASS('a, equal_class) \<equiv> (((=) :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> equal)" |
63575 | 2037 |
(is "?ofclass \<equiv> ?equal") |
31956
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2038 |
proof |
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2039 |
assume "PROP ?ofclass" |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
2040 |
show "PROP ?equal" |
69593 | 2041 |
by (tactic \<open>ALLGOALS (resolve_tac \<^context> [Thm.unconstrainT @{thm eq_equal}])\<close>) |
60758 | 2042 |
(fact \<open>PROP ?ofclass\<close>) |
31956
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2043 |
next |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
2044 |
assume "PROP ?equal" |
31956
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2045 |
show "PROP ?ofclass" proof |
60758 | 2046 |
qed (simp add: \<open>PROP ?equal\<close>) |
31956
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2047 |
qed |
c3844c4d0c2c
more accurate certificates for constant aliasses
haftmann
parents:
31902
diff
changeset
|
2048 |
|
69593 | 2049 |
setup \<open>Sign.add_const_constraint (\<^const_name>\<open>equal\<close>, SOME \<^typ>\<open>'a::equal \<Rightarrow> 'a \<Rightarrow> bool\<close>)\<close> |
58826 | 2050 |
|
60758 | 2051 |
setup \<open>Nbe.add_const_alias @{thm equal_alias_cert}\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2052 |
|
60758 | 2053 |
text \<open>Cases\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2054 |
|
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2055 |
lemma Let_case_cert: |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2056 |
assumes "CASE \<equiv> (\<lambda>x. Let x f)" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2057 |
shows "CASE x \<equiv> f x" |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2058 |
using assms by simp_all |
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2059 |
|
60758 | 2060 |
setup \<open> |
66251
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
haftmann
parents:
66109
diff
changeset
|
2061 |
Code.declare_case_global @{thm Let_case_cert} #> |
69593 | 2062 |
Code.declare_undefined_global \<^const_name>\<open>undefined\<close> |
60758 | 2063 |
\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2064 |
|
54890
cb892d835803
fundamental treatment of undefined vs. universally partial replaces code_abort
haftmann
parents:
54742
diff
changeset
|
2065 |
declare [[code abort: undefined]] |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2066 |
|
38972 | 2067 |
|
60758 | 2068 |
subsubsection \<open>Generic code generator target languages\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2069 |
|
69593 | 2070 |
text \<open>type \<^typ>\<open>bool\<close>\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2071 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2072 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2073 |
type_constructor bool \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2074 |
(SML) "bool" and (OCaml) "bool" and (Haskell) "Bool" and (Scala) "Boolean" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2075 |
| constant True \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2076 |
(SML) "true" and (OCaml) "true" and (Haskell) "True" and (Scala) "true" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2077 |
| constant False \<rightharpoonup> |
58826 | 2078 |
(SML) "false" and (OCaml) "false" and (Haskell) "False" and (Scala) "false" |
34294 | 2079 |
|
81706 | 2080 |
code_reserved |
2081 |
(SML) bool true false |
|
2082 |
and (OCaml) bool |
|
2083 |
and (Scala) Boolean |
|
34294 | 2084 |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2085 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2086 |
constant Not \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2087 |
(SML) "not" and (OCaml) "not" and (Haskell) "not" and (Scala) "'! _" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2088 |
| constant HOL.conj \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2089 |
(SML) infixl 1 "andalso" and (OCaml) infixl 3 "&&" and (Haskell) infixr 3 "&&" and (Scala) infixl 3 "&&" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2090 |
| constant HOL.disj \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2091 |
(SML) infixl 0 "orelse" and (OCaml) infixl 2 "||" and (Haskell) infixl 2 "||" and (Scala) infixl 1 "||" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2092 |
| constant HOL.implies \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2093 |
(SML) "!(if (_)/ then (_)/ else true)" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2094 |
and (OCaml) "!(if (_)/ then (_)/ else true)" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2095 |
and (Haskell) "!(if (_)/ then (_)/ else True)" |
80088
5afbf04418ec
avoid Scala if-expressions and thus make it work both for -new-syntax or -old-syntax;
wenzelm
parents:
79772
diff
changeset
|
2096 |
and (Scala) "!((_) match {/ case true => (_)/ case false => true/ })" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2097 |
| constant If \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2098 |
(SML) "!(if (_)/ then (_)/ else (_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2099 |
and (OCaml) "!(if (_)/ then (_)/ else (_))" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2100 |
and (Haskell) "!(if (_)/ then (_)/ else (_))" |
80088
5afbf04418ec
avoid Scala if-expressions and thus make it work both for -new-syntax or -old-syntax;
wenzelm
parents:
79772
diff
changeset
|
2101 |
and (Scala) "!((_) match {/ case true => (_)/ case false => (_)/ })" |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2102 |
|
81706 | 2103 |
code_reserved |
2104 |
(SML) not |
|
2105 |
and (OCaml) not |
|
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2106 |
|
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2107 |
code_identifier |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2108 |
code_module Pure \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2109 |
(SML) HOL and (OCaml) HOL and (Haskell) HOL and (Scala) HOL |
39026 | 2110 |
|
63575 | 2111 |
text \<open>Using built-in Haskell equality.\<close> |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2112 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2113 |
type_class equal \<rightharpoonup> (Haskell) "Eq" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2114 |
| constant HOL.equal \<rightharpoonup> (Haskell) infix 4 "==" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2115 |
| constant HOL.eq \<rightharpoonup> (Haskell) infix 4 "==" |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2116 |
|
63575 | 2117 |
text \<open>\<open>undefined\<close>\<close> |
52435
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2118 |
code_printing |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2119 |
constant undefined \<rightharpoonup> |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2120 |
(SML) "!(raise/ Fail/ \"undefined\")" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2121 |
and (OCaml) "failwith/ \"undefined\"" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2122 |
and (Haskell) "error/ \"undefined\"" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2123 |
and (Scala) "!sys.error(\"undefined\")" |
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents:
52432
diff
changeset
|
2124 |
|
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2125 |
|
60758 | 2126 |
subsubsection \<open>Evaluation and normalization by evaluation\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2127 |
|
60758 | 2128 |
method_setup eval = \<open> |
58826 | 2129 |
let |
2130 |
fun eval_tac ctxt = |
|
74536 | 2131 |
let val conv = Code_Runtime.dynamic_holds_conv |
58839 | 2132 |
in |
74536 | 2133 |
CONVERSION (Conv.params_conv ~1 (Conv.concl_conv ~1 o conv) ctxt) THEN' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59028
diff
changeset
|
2134 |
resolve_tac ctxt [TrueI] |
58839 | 2135 |
end |
58826 | 2136 |
in |
2137 |
Scan.succeed (SIMPLE_METHOD' o eval_tac) |
|
2138 |
end |
|
60758 | 2139 |
\<close> "solve goal by evaluation" |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2140 |
|
60758 | 2141 |
method_setup normalization = \<open> |
46190
a42c5f23109f
more conventional eval_tac vs. method_setup "eval";
wenzelm
parents:
46161
diff
changeset
|
2142 |
Scan.succeed (fn ctxt => |
a42c5f23109f
more conventional eval_tac vs. method_setup "eval";
wenzelm
parents:
46161
diff
changeset
|
2143 |
SIMPLE_METHOD' |
a42c5f23109f
more conventional eval_tac vs. method_setup "eval";
wenzelm
parents:
46161
diff
changeset
|
2144 |
(CHANGED_PROP o |
55757 | 2145 |
(CONVERSION (Nbe.dynamic_conv ctxt) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59028
diff
changeset
|
2146 |
THEN_ALL_NEW (TRY o resolve_tac ctxt [TrueI])))) |
60758 | 2147 |
\<close> "solve goal by normalization" |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2148 |
|
31902 | 2149 |
|
60758 | 2150 |
subsection \<open>Counterexample Search Units\<close> |
33084 | 2151 |
|
60758 | 2152 |
subsubsection \<open>Quickcheck\<close> |
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2153 |
|
33084 | 2154 |
quickcheck_params [size = 5, iterations = 50] |
2155 |
||
30929
d9343c0aac11
code generator bootstrap theory src/Tools/Code_Generator.thy
haftmann
parents:
30927
diff
changeset
|
2156 |
|
60758 | 2157 |
subsubsection \<open>Nitpick setup\<close> |
30309
188f0658af9f
Added a "nitpick_maybe" symbol, which is used by Nitpick. This will go away once Nitpick is part of HOL.
blanchet
parents:
30254
diff
changeset
|
2158 |
|
59028 | 2159 |
named_theorems nitpick_unfold "alternative definitions of constants as needed by Nitpick" |
2160 |
and nitpick_simp "equational specification of constants as needed by Nitpick" |
|
2161 |
and nitpick_psimp "partial equational specification of constants as needed by Nitpick" |
|
2162 |
and nitpick_choice_spec "choice specification of constants as needed by Nitpick" |
|
30980 | 2163 |
|
41792
ff3cb0c418b7
renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents:
41636
diff
changeset
|
2164 |
declare if_bool_eq_conj [nitpick_unfold, no_atp] |
63575 | 2165 |
and if_bool_eq_disj [no_atp] |
41792
ff3cb0c418b7
renamed "nitpick\_def" to "nitpick_unfold" to reflect its new semantics
blanchet
parents:
41636
diff
changeset
|
2166 |
|
29863
dadad1831e9d
Added "nitpick_const_simps" and "nitpick_ind_intros" attributes for theorems;
blanchet
parents:
29608
diff
changeset
|
2167 |
|
60758 | 2168 |
subsection \<open>Preprocessing for the predicate compiler\<close> |
33084 | 2169 |
|
59028 | 2170 |
named_theorems code_pred_def "alternative definitions of constants for the Predicate Compiler" |
2171 |
and code_pred_inline "inlining definitions for the Predicate Compiler" |
|
2172 |
and code_pred_simp "simplification rules for the optimisations in the Predicate Compiler" |
|
33084 | 2173 |
|
2174 |
||
60758 | 2175 |
subsection \<open>Legacy tactics and ML bindings\<close> |
21671 | 2176 |
|
60758 | 2177 |
ML \<open> |
58826 | 2178 |
(* combination of (spec RS spec RS ...(j times) ... spec RS mp) *) |
2179 |
local |
|
80662 | 2180 |
fun wrong_prem \<^Const_>\<open>All _ for \<open>Abs (_, _, t)\<close>\<close> = wrong_prem t |
58826 | 2181 |
| wrong_prem (Bound _) = true |
2182 |
| wrong_prem _ = false; |
|
81954
6f2bcdfa9a19
misc tuning: more concise operations on prems (without change of exceptions);
wenzelm
parents:
81706
diff
changeset
|
2183 |
val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.take_prems_of 1); |
61914 | 2184 |
fun smp i = funpow i (fn m => filter_right ([spec] RL m)) [mp]; |
58826 | 2185 |
in |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59028
diff
changeset
|
2186 |
fun smp_tac ctxt j = EVERY' [dresolve_tac ctxt (smp j), assume_tac ctxt]; |
58826 | 2187 |
end; |
22839 | 2188 |
|
58826 | 2189 |
local |
2190 |
val nnf_ss = |
|
69593 | 2191 |
simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms nnf_simps}); |
58826 | 2192 |
in |
2193 |
fun nnf_conv ctxt = Simplifier.rewrite (put_simpset nnf_ss ctxt); |
|
2194 |
end |
|
60758 | 2195 |
\<close> |
21671 | 2196 |
|
38866 | 2197 |
hide_const (open) eq equal |
2198 |
||
14357 | 2199 |
end |