| author | nipkow | 
| Wed, 09 Jul 2014 11:48:21 +0200 | |
| changeset 57537 | 810bc6c41ebd | 
| parent 57233 | 8fcbfce2a2a9 | 
| child 57983 | 6edc3529bb4e | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Product_Type.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1992 University of Cambridge | |
| 11777 | 4 | *) | 
| 10213 | 5 | |
| 11838 | 6 | header {* Cartesian products *}
 | 
| 10213 | 7 | |
| 15131 | 8 | theory Product_Type | 
| 33959 
2afc55e8ed27
bootstrap datatype_rep_proofs in Datatype.thy (avoids unchecked dynamic name references)
 haftmann parents: 
33638diff
changeset | 9 | imports Typedef Inductive Fun | 
| 46950 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 wenzelm parents: 
46630diff
changeset | 10 | keywords "inductive_set" "coinductive_set" :: thy_decl | 
| 15131 | 11 | begin | 
| 11838 | 12 | |
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 13 | subsection {* @{typ bool} is a datatype *}
 | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 14 | |
| 57091 | 15 | free_constructors case_bool for True | False | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 16 | by auto | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 17 | |
| 55442 | 18 | text {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
 | 
| 19 | ||
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 20 | setup {* Sign.mandatory_path "old" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 21 | |
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26975diff
changeset | 22 | rep_datatype True False by (auto intro: bool_induct) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 23 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 24 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 25 | |
| 55468 
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
 blanchet parents: 
55467diff
changeset | 26 | text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
 | 
| 55442 | 27 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 28 | setup {* Sign.mandatory_path "bool" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 29 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 30 | lemmas induct = old.bool.induct | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 31 | lemmas inducts = old.bool.inducts | 
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 32 | lemmas rec = old.bool.rec | 
| 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 33 | lemmas simps = bool.distinct bool.case bool.rec | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 34 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 35 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 36 | |
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 37 | declare case_split [cases type: bool] | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 38 | -- "prefer plain propositional version" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 39 | |
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28262diff
changeset | 40 | lemma | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38715diff
changeset | 41 | shows [code]: "HOL.equal False P \<longleftrightarrow> \<not> P" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38715diff
changeset | 42 | and [code]: "HOL.equal True P \<longleftrightarrow> P" | 
| 46630 | 43 | and [code]: "HOL.equal P False \<longleftrightarrow> \<not> P" | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38715diff
changeset | 44 | and [code]: "HOL.equal P True \<longleftrightarrow> P" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38715diff
changeset | 45 | and [code nbe]: "HOL.equal P P \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38715diff
changeset | 46 | by (simp_all add: equal) | 
| 25534 
d0b74fdd6067
simplified infrastructure for code generator operational equality
 haftmann parents: 
25511diff
changeset | 47 | |
| 43654 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 48 | lemma If_case_cert: | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 49 | assumes "CASE \<equiv> (\<lambda>b. If b f g)" | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 50 | shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)" | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 51 | using assms by simp_all | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 52 | |
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 53 | setup {*
 | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 54 |   Code.add_case @{thm If_case_cert}
 | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 55 | *} | 
| 
3f1a44c2d645
install case certificate for If after code_datatype declaration for bool
 haftmann parents: 
43595diff
changeset | 56 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 57 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 58 | constant "HOL.equal :: bool \<Rightarrow> bool \<Rightarrow> bool" \<rightharpoonup> (Haskell) infix 4 "==" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 59 | | class_instance "bool" :: "equal" \<rightharpoonup> (Haskell) - | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 60 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 61 | |
| 37166 | 62 | subsection {* The @{text unit} type *}
 | 
| 11838 | 63 | |
| 49834 | 64 | typedef unit = "{True}"
 | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45662diff
changeset | 65 | by auto | 
| 11838 | 66 | |
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45662diff
changeset | 67 | definition Unity :: unit  ("'(')")
 | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45662diff
changeset | 68 | where "() = Abs_unit True" | 
| 11838 | 69 | |
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 70 | lemma unit_eq [no_atp]: "u = ()" | 
| 40590 | 71 | by (induct u) (simp add: Unity_def) | 
| 11838 | 72 | |
| 73 | text {*
 | |
| 74 |   Simplification procedure for @{thm [source] unit_eq}.  Cannot use
 | |
| 75 | this rule directly --- it loops! | |
| 76 | *} | |
| 77 | ||
| 43594 | 78 | simproc_setup unit_eq ("x::unit") = {*
 | 
| 79 | fn _ => fn _ => fn ct => | |
| 80 | if HOLogic.is_unit (term_of ct) then NONE | |
| 81 |     else SOME (mk_meta_eq @{thm unit_eq})
 | |
| 11838 | 82 | *} | 
| 83 | ||
| 55469 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
55468diff
changeset | 84 | free_constructors case_unit for "()" | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 85 | by auto | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 86 | |
| 55442 | 87 | text {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
 | 
| 88 | ||
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 89 | setup {* Sign.mandatory_path "old" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 90 | |
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26975diff
changeset | 91 | rep_datatype "()" by simp | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 92 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 93 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 94 | |
| 55468 
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
 blanchet parents: 
55467diff
changeset | 95 | text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
 | 
| 55442 | 96 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 97 | setup {* Sign.mandatory_path "unit" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 98 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 99 | lemmas induct = old.unit.induct | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 100 | lemmas inducts = old.unit.inducts | 
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 101 | lemmas rec = old.unit.rec | 
| 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 102 | lemmas simps = unit.case unit.rec | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 103 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 104 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 105 | |
| 11838 | 106 | lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()" | 
| 107 | by simp | |
| 108 | ||
| 109 | lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P" | |
| 110 | by (rule triv_forall_equality) | |
| 111 | ||
| 112 | text {*
 | |
| 43594 | 113 |   This rewrite counters the effect of simproc @{text unit_eq} on @{term
 | 
| 11838 | 114 |   [source] "%u::unit. f u"}, replacing it by @{term [source]
 | 
| 115 |   f} rather than by @{term [source] "%u. f ()"}.
 | |
| 116 | *} | |
| 117 | ||
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
52435diff
changeset | 118 | lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f" | 
| 11838 | 119 | by (rule ext) simp | 
| 10213 | 120 | |
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
52435diff
changeset | 121 | lemma UNIV_unit: | 
| 43866 
8a50dc70cbff
moving UNIV = ... equations to their proper theories
 haftmann parents: 
43654diff
changeset | 122 |   "UNIV = {()}" by auto
 | 
| 
8a50dc70cbff
moving UNIV = ... equations to their proper theories
 haftmann parents: 
43654diff
changeset | 123 | |
| 30924 | 124 | instantiation unit :: default | 
| 125 | begin | |
| 126 | ||
| 127 | definition "default = ()" | |
| 128 | ||
| 129 | instance .. | |
| 130 | ||
| 131 | end | |
| 10213 | 132 | |
| 57233 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 133 | instantiation unit :: "{complete_boolean_algebra, complete_linorder, wellorder}"
 | 
| 57016 | 134 | begin | 
| 135 | ||
| 57233 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 136 | definition less_eq_unit :: "unit \<Rightarrow> unit \<Rightarrow> bool" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 137 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 138 | "(_::unit) \<le> _ \<longleftrightarrow> True" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 139 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 140 | lemma less_eq_unit [iff]: | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 141 | "(u::unit) \<le> v" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 142 | by (simp add: less_eq_unit_def) | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 143 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 144 | definition less_unit :: "unit \<Rightarrow> unit \<Rightarrow> bool" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 145 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 146 | "(_::unit) < _ \<longleftrightarrow> False" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 147 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 148 | lemma less_unit [iff]: | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 149 | "\<not> (u::unit) < v" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 150 | by (simp_all add: less_eq_unit_def less_unit_def) | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 151 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 152 | definition bot_unit :: unit | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 153 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 154 | [code_unfold]: "\<bottom> = ()" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 155 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 156 | definition top_unit :: unit | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 157 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 158 | [code_unfold]: "\<top> = ()" | 
| 57016 | 159 | |
| 57233 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 160 | definition inf_unit :: "unit \<Rightarrow> unit \<Rightarrow> unit" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 161 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 162 | [simp]: "_ \<sqinter> _ = ()" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 163 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 164 | definition sup_unit :: "unit \<Rightarrow> unit \<Rightarrow> unit" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 165 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 166 | [simp]: "_ \<squnion> _ = ()" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 167 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 168 | definition Inf_unit :: "unit set \<Rightarrow> unit" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 169 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 170 | [simp]: "\<Sqinter>_ = ()" | 
| 57016 | 171 | |
| 57233 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 172 | definition Sup_unit :: "unit set \<Rightarrow> unit" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 173 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 174 | [simp]: "\<Squnion>_ = ()" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 175 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 176 | definition uminus_unit :: "unit \<Rightarrow> unit" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 177 | where | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 178 | [simp]: "- _ = ()" | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 179 | |
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 180 | declare less_eq_unit_def [abs_def, code_unfold] | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 181 | less_unit_def [abs_def, code_unfold] | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 182 | inf_unit_def [abs_def, code_unfold] | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 183 | sup_unit_def [abs_def, code_unfold] | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 184 | Inf_unit_def [abs_def, code_unfold] | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 185 | Sup_unit_def [abs_def, code_unfold] | 
| 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 186 | uminus_unit_def [abs_def, code_unfold] | 
| 57016 | 187 | |
| 188 | instance | |
| 57233 
8fcbfce2a2a9
uniform treatment of trivial unit instances: simplify by default, unfold in code preprocessor
 haftmann parents: 
57201diff
changeset | 189 | by intro_classes auto | 
| 57016 | 190 | |
| 191 | end | |
| 192 | ||
| 28562 | 193 | lemma [code]: | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38715diff
changeset | 194 | "HOL.equal (u\<Colon>unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+ | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 195 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 196 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 197 | type_constructor unit \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 198 | (SML) "unit" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 199 | and (OCaml) "unit" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 200 | and (Haskell) "()" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 201 | and (Scala) "Unit" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 202 | | constant Unity \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 203 | (SML) "()" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 204 | and (OCaml) "()" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 205 | and (Haskell) "()" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 206 | and (Scala) "()" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 207 | | class_instance unit :: equal \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 208 | (Haskell) - | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 209 | | constant "HOL.equal :: unit \<Rightarrow> unit \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 210 | (Haskell) infix 4 "==" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 211 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 212 | code_reserved SML | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 213 | unit | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 214 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 215 | code_reserved OCaml | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 216 | unit | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 217 | |
| 34886 | 218 | code_reserved Scala | 
| 219 | Unit | |
| 220 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 221 | |
| 37166 | 222 | subsection {* The product type *}
 | 
| 10213 | 223 | |
| 37166 | 224 | subsubsection {* Type definition *}
 | 
| 225 | ||
| 226 | definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 227 | "Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)" | 
| 10213 | 228 | |
| 45696 | 229 | definition "prod = {f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
 | 
| 230 | ||
| 49834 | 231 | typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \<Rightarrow> 'b \<Rightarrow> bool) set"
 | 
| 45696 | 232 | unfolding prod_def by auto | 
| 10213 | 233 | |
| 35427 | 234 | type_notation (xsymbols) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 235 |   "prod"  ("(_ \<times>/ _)" [21, 20] 20)
 | 
| 35427 | 236 | type_notation (HTML output) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 237 |   "prod"  ("(_ \<times>/ _)" [21, 20] 20)
 | 
| 10213 | 238 | |
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 239 | definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b" where | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 240 | "Pair a b = Abs_prod (Pair_Rep a b)" | 
| 37166 | 241 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 242 | lemma prod_cases: "(\<And>a b. P (Pair a b)) \<Longrightarrow> P p" | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 243 | by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 244 | |
| 55469 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
55468diff
changeset | 245 | free_constructors case_prod for Pair fst snd | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 246 | proof - | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 247 | fix P :: bool and p :: "'a \<times> 'b" | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 248 | show "(\<And>x1 x2. p = Pair x1 x2 \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 249 | by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) | 
| 37166 | 250 | next | 
| 251 | fix a c :: 'a and b d :: 'b | |
| 252 | have "Pair_Rep a b = Pair_Rep c d \<longleftrightarrow> a = c \<and> b = d" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 253 | by (auto simp add: Pair_Rep_def fun_eq_iff) | 
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 254 | moreover have "Pair_Rep a b \<in> prod" and "Pair_Rep c d \<in> prod" | 
| 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 255 | by (auto simp add: prod_def) | 
| 37166 | 256 | ultimately show "Pair a b = Pair c d \<longleftrightarrow> a = c \<and> b = d" | 
| 37389 
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
 haftmann parents: 
37387diff
changeset | 257 | by (simp add: Pair_def Abs_prod_inject) | 
| 37166 | 258 | qed | 
| 259 | ||
| 55442 | 260 | text {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
 | 
| 261 | ||
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 262 | setup {* Sign.mandatory_path "old" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 263 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 264 | rep_datatype Pair | 
| 55403 
677569668824
avoid duplicate 'case' definitions by first looking up 'Ctr_Sugar'
 blanchet parents: 
55393diff
changeset | 265 | by (erule prod_cases) (rule prod.inject) | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 266 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 267 | setup {* Sign.parent_path *}
 | 
| 37704 
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
 blanchet parents: 
37678diff
changeset | 268 | |
| 55468 
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
 blanchet parents: 
55467diff
changeset | 269 | text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
 | 
| 55442 | 270 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 271 | setup {* Sign.mandatory_path "prod" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 272 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 273 | declare | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 274 | old.prod.inject[iff del] | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 275 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 276 | lemmas induct = old.prod.induct | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 277 | lemmas inducts = old.prod.inducts | 
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 278 | lemmas rec = old.prod.rec | 
| 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 279 | lemmas simps = prod.inject prod.case prod.rec | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 280 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 281 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 282 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 283 | declare prod.case [nitpick_simp del] | 
| 40929 
7ff03a5e044f
theorem names generated by the (rep_)datatype command now have mandatory qualifiers
 huffman parents: 
40702diff
changeset | 284 | declare prod.weak_case_cong [cong del] | 
| 37411 
c88c44156083
removed simplifier congruence rule of "prod_case"
 haftmann parents: 
37389diff
changeset | 285 | |
| 37166 | 286 | |
| 287 | subsubsection {* Tuple syntax *}
 | |
| 288 | ||
| 37591 | 289 | abbreviation (input) split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
 | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 290 | "split \<equiv> case_prod" | 
| 19535 | 291 | |
| 11777 | 292 | text {*
 | 
| 293 |   Patterns -- extends pre-defined type @{typ pttrn} used in
 | |
| 294 | abstractions. | |
| 295 | *} | |
| 10213 | 296 | |
| 41229 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
40968diff
changeset | 297 | nonterminal tuple_args and patterns | 
| 10213 | 298 | |
| 299 | syntax | |
| 300 |   "_tuple"      :: "'a => tuple_args => 'a * 'b"        ("(1'(_,/ _'))")
 | |
| 301 |   "_tuple_arg"  :: "'a => tuple_args"                   ("_")
 | |
| 302 |   "_tuple_args" :: "'a => tuple_args => tuple_args"     ("_,/ _")
 | |
| 11025 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 303 |   "_pattern"    :: "[pttrn, patterns] => pttrn"         ("'(_,/ _')")
 | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 304 |   ""            :: "pttrn => patterns"                  ("_")
 | 
| 
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
 oheimb parents: 
10289diff
changeset | 305 |   "_patterns"   :: "[pttrn, patterns] => patterns"      ("_,/ _")
 | 
| 10213 | 306 | |
| 307 | translations | |
| 35115 | 308 | "(x, y)" == "CONST Pair x y" | 
| 51392 
635562bc14ef
extended set comprehension notation with {pttrn : A . P}
 nipkow parents: 
51173diff
changeset | 309 | "_pattern x y" => "CONST Pair x y" | 
| 
635562bc14ef
extended set comprehension notation with {pttrn : A . P}
 nipkow parents: 
51173diff
changeset | 310 | "_patterns x y" => "CONST Pair x y" | 
| 10213 | 311 | "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 312 | "%(x, y, zs). b" == "CONST case_prod (%x (y, zs). b)" | 
| 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 313 | "%(x, y). b" == "CONST case_prod (%x y. b)" | 
| 35115 | 314 | "_abs (CONST Pair x y) t" => "%(x, y). t" | 
| 37166 | 315 |   -- {* The last rule accommodates tuples in `case C ... (x,y) ... => ...'
 | 
| 316 | The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *} | |
| 10213 | 317 | |
| 35115 | 318 | (*reconstruct pattern from (nested) splits, avoiding eta-contraction of body; | 
| 319 | works best with enclosing "let", if "let" does not avoid eta-contraction*) | |
| 14359 | 320 | print_translation {*
 | 
| 52143 | 321 | let | 
| 322 | fun split_tr' [Abs (x, T, t as (Abs abs))] = | |
| 323 | (* split (%x y. t) => %(x,y) t *) | |
| 324 | let | |
| 325 | val (y, t') = Syntax_Trans.atomic_abs_tr' abs; | |
| 326 | val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); | |
| 327 | in | |
| 328 |             Syntax.const @{syntax_const "_abs"} $
 | |
| 329 |               (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
 | |
| 330 | end | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 331 |       | split_tr' [Abs (x, T, (s as Const (@{const_syntax case_prod}, _) $ t))] =
 | 
| 52143 | 332 | (* split (%x. (split (%y z. t))) => %(x,y,z). t *) | 
| 333 | let | |
| 334 |             val Const (@{syntax_const "_abs"}, _) $
 | |
| 335 |               (Const (@{syntax_const "_pattern"}, _) $ y $ z) $ t' = split_tr' [t];
 | |
| 336 | val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); | |
| 337 | in | |
| 338 |             Syntax.const @{syntax_const "_abs"} $
 | |
| 339 |               (Syntax.const @{syntax_const "_pattern"} $ x' $
 | |
| 340 |                 (Syntax.const @{syntax_const "_patterns"} $ y $ z)) $ t''
 | |
| 341 | end | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 342 |       | split_tr' [Const (@{const_syntax case_prod}, _) $ t] =
 | 
| 52143 | 343 | (* split (split (%x y z. t)) => %((x, y), z). t *) | 
| 344 | split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *) | |
| 345 |       | split_tr' [Const (@{syntax_const "_abs"}, _) $ x_y $ Abs abs] =
 | |
| 346 | (* split (%pttrn z. t) => %(pttrn,z). t *) | |
| 347 | let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in | |
| 348 |             Syntax.const @{syntax_const "_abs"} $
 | |
| 349 |               (Syntax.const @{syntax_const "_pattern"} $ x_y $ z) $ t
 | |
| 350 | end | |
| 351 | | split_tr' _ = raise Match; | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 352 |   in [(@{const_syntax case_prod}, K split_tr')] end
 | 
| 14359 | 353 | *} | 
| 354 | ||
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 355 | (* print "split f" as "\<lambda>(x,y). f x y" and "split (\<lambda>x. f x)" as "\<lambda>(x,y). f x y" *) | 
| 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 356 | typed_print_translation {*
 | 
| 52143 | 357 | let | 
| 358 | fun split_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match | |
| 359 | | split_guess_names_tr' T [Abs (x, xT, t)] = | |
| 360 | (case (head_of t) of | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 361 |             Const (@{const_syntax case_prod}, _) => raise Match
 | 
| 52143 | 362 | | _ => | 
| 363 | let | |
| 364 | val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; | |
| 365 |               val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0);
 | |
| 366 | val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t'); | |
| 367 | in | |
| 368 |               Syntax.const @{syntax_const "_abs"} $
 | |
| 369 |                 (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
 | |
| 370 | end) | |
| 371 | | split_guess_names_tr' T [t] = | |
| 372 | (case head_of t of | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 373 |             Const (@{const_syntax case_prod}, _) => raise Match
 | 
| 52143 | 374 | | _ => | 
| 375 | let | |
| 376 | val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; | |
| 377 | val (y, t') = | |
| 378 |                 Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0);
 | |
| 379 |               val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t');
 | |
| 380 | in | |
| 381 |               Syntax.const @{syntax_const "_abs"} $
 | |
| 382 |                 (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t''
 | |
| 383 | end) | |
| 384 | | split_guess_names_tr' _ _ = raise Match; | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 385 |   in [(@{const_syntax case_prod}, K split_guess_names_tr')] end
 | 
| 15422 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 386 | *} | 
| 
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
 schirmer parents: 
15404diff
changeset | 387 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 388 | (* Force eta-contraction for terms of the form "Q A (%p. case_prod P p)" | 
| 42059 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 389 | where Q is some bounded quantifier or set operator. | 
| 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 390 | Reason: the above prints as "Q p : A. case p of (x,y) \<Rightarrow> P x y" | 
| 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 391 | whereas we want "Q (x,y):A. P x y". | 
| 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 392 | Otherwise prevent eta-contraction. | 
| 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 393 | *) | 
| 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 394 | print_translation {*
 | 
| 52143 | 395 | let | 
| 396 | fun contract Q tr ctxt ts = | |
| 397 | (case ts of | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 398 |         [A, Abs (_, _, (s as Const (@{const_syntax case_prod},_) $ t) $ Bound 0)] =>
 | 
| 52143 | 399 | if Term.is_dependent t then tr ctxt ts | 
| 400 | else Syntax.const Q $ A $ s | |
| 401 | | _ => tr ctxt ts); | |
| 402 | in | |
| 42284 | 403 |     [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Ball} @{syntax_const "_Ball"},
 | 
| 404 |      Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Bex} @{syntax_const "_Bex"},
 | |
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56077diff
changeset | 405 |      Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFIMUM} @{syntax_const "_INF"},
 | 
| 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56077diff
changeset | 406 |      Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPREMUM} @{syntax_const "_SUP"}]
 | 
| 52143 | 407 | |> map (fn (Q, tr) => (Q, contract Q tr)) | 
| 408 | end | |
| 42059 
83f3dc509068
fixed a printing problem for bounded quantifiers and bounded set operators in the case of tuples
 nipkow parents: 
41792diff
changeset | 409 | *} | 
| 10213 | 410 | |
| 37166 | 411 | subsubsection {* Code generator setup *}
 | 
| 412 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 413 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 414 | type_constructor prod \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 415 | (SML) infix 2 "*" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 416 | and (OCaml) infix 2 "*" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 417 | and (Haskell) "!((_),/ (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 418 | and (Scala) "((_),/ (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 419 | | constant Pair \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 420 | (SML) "!((_),/ (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 421 | and (OCaml) "!((_),/ (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 422 | and (Haskell) "!((_),/ (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 423 | and (Scala) "!((_),/ (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 424 | | class_instance prod :: equal \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 425 | (Haskell) - | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 426 | | constant "HOL.equal :: 'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 427 | (Haskell) infix 4 "==" | 
| 37166 | 428 | |
| 429 | ||
| 430 | subsubsection {* Fundamental operations and properties *}
 | |
| 11838 | 431 | |
| 49897 
cc69be3c8f87
moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
 bulwahn parents: 
49834diff
changeset | 432 | lemma Pair_inject: | 
| 
cc69be3c8f87
moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
 bulwahn parents: 
49834diff
changeset | 433 | assumes "(a, b) = (a', b')" | 
| 
cc69be3c8f87
moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
 bulwahn parents: 
49834diff
changeset | 434 | and "a = a' ==> b = b' ==> R" | 
| 
cc69be3c8f87
moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
 bulwahn parents: 
49834diff
changeset | 435 | shows R | 
| 
cc69be3c8f87
moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
 bulwahn parents: 
49834diff
changeset | 436 | using assms by simp | 
| 
cc69be3c8f87
moving Pair_inject from legacy and duplicate section to general section, as Pair_inject was considered a duplicate in e8400e31528a by mistake (cf. communication on dev mailing list)
 bulwahn parents: 
49834diff
changeset | 437 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 438 | lemma surj_pair [simp]: "EX x y. p = (x, y)" | 
| 37166 | 439 | by (cases p) simp | 
| 10213 | 440 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 441 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 442 | constant fst \<rightharpoonup> (Haskell) "fst" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 443 | | constant snd \<rightharpoonup> (Haskell) "snd" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 444 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 445 | lemma case_prod_unfold [nitpick_unfold]: "case_prod = (%c p. c (fst p) (snd p))" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 446 | by (simp add: fun_eq_iff split: prod.split) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 447 | |
| 11838 | 448 | lemma fst_eqD: "fst (x, y) = a ==> x = a" | 
| 449 | by simp | |
| 450 | ||
| 451 | lemma snd_eqD: "snd (x, y) = a ==> y = a" | |
| 452 | by simp | |
| 453 | ||
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 454 | lemmas surjective_pairing = prod.collapse [symmetric] | 
| 11838 | 455 | |
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
43866diff
changeset | 456 | lemma prod_eq_iff: "s = t \<longleftrightarrow> fst s = fst t \<and> snd s = snd t" | 
| 37166 | 457 | by (cases s, cases t) simp | 
| 458 | ||
| 459 | lemma prod_eqI [intro?]: "fst p = fst q \<Longrightarrow> snd p = snd q \<Longrightarrow> p = q" | |
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
43866diff
changeset | 460 | by (simp add: prod_eq_iff) | 
| 37166 | 461 | |
| 462 | lemma split_conv [simp, code]: "split f (a, b) = f a b" | |
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 463 | by (fact prod.case) | 
| 37166 | 464 | |
| 465 | lemma splitI: "f a b \<Longrightarrow> split f (a, b)" | |
| 466 | by (rule split_conv [THEN iffD2]) | |
| 467 | ||
| 468 | lemma splitD: "split f (a, b) \<Longrightarrow> f a b" | |
| 469 | by (rule split_conv [THEN iffD1]) | |
| 470 | ||
| 471 | lemma split_Pair [simp]: "(\<lambda>(x, y). (x, y)) = id" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 472 | by (simp add: fun_eq_iff split: prod.split) | 
| 37166 | 473 | |
| 474 | lemma split_eta: "(\<lambda>(x, y). f (x, y)) = f" | |
| 475 |   -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *}
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 476 | by (simp add: fun_eq_iff split: prod.split) | 
| 37166 | 477 | |
| 478 | lemma split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)" | |
| 479 | by (cases x) simp | |
| 480 | ||
| 481 | lemma split_twice: "split f (split g p) = split (\<lambda>x y. split f (g x y)) p" | |
| 482 | by (cases p) simp | |
| 483 | ||
| 484 | lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))" | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 485 | by (simp add: case_prod_unfold) | 
| 37166 | 486 | |
| 487 | lemma split_weak_cong: "p = q \<Longrightarrow> split c p = split c q" | |
| 488 |   -- {* Prevents simplification of @{term c}: much faster *}
 | |
| 40929 
7ff03a5e044f
theorem names generated by the (rep_)datatype command now have mandatory qualifiers
 huffman parents: 
40702diff
changeset | 489 | by (fact prod.weak_case_cong) | 
| 37166 | 490 | |
| 491 | lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" | |
| 492 | by (simp add: split_eta) | |
| 493 | ||
| 47740 
a8989fe9a3a5
added "no_atp"s for extremely prolific, useless facts for ATPs
 blanchet parents: 
46950diff
changeset | 494 | lemma split_paired_all [no_atp]: "(!!x. PROP P x) == (!!a b. PROP P (a, b))" | 
| 11820 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 495 | proof | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 496 | fix a b | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 497 | assume "!!x. PROP P x" | 
| 19535 | 498 | then show "PROP P (a, b)" . | 
| 11820 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 499 | next | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 500 | fix x | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 501 | assume "!!a b. PROP P (a, b)" | 
| 19535 | 502 | from `PROP P (fst x, snd x)` show "PROP P x" by simp | 
| 11820 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 503 | qed | 
| 
015a82d4ee96
proper proof of split_paired_all (presently unused);
 wenzelm parents: 
11777diff
changeset | 504 | |
| 50104 | 505 | lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))" | 
| 506 | by (cases x) simp | |
| 507 | ||
| 11838 | 508 | text {*
 | 
| 509 |   The rule @{thm [source] split_paired_all} does not work with the
 | |
| 510 | Simplifier because it also affects premises in congrence rules, | |
| 511 |   where this can lead to premises of the form @{text "!!a b. ... =
 | |
| 512 | ?P(a, b)"} which cannot be solved by reflexivity. | |
| 513 | *} | |
| 514 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 515 | lemmas split_tupled_all = split_paired_all unit_all_eq2 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 516 | |
| 26480 | 517 | ML {*
 | 
| 11838 | 518 | (* replace parameters of product type by individual component parameters *) | 
| 519 | local (* filtering with exists_paired_all is an essential optimization *) | |
| 56245 | 520 |     fun exists_paired_all (Const (@{const_name Pure.all}, _) $ Abs (_, T, t)) =
 | 
| 11838 | 521 | can HOLogic.dest_prodT T orelse exists_paired_all t | 
| 522 | | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u | |
| 523 | | exists_paired_all (Abs (_, _, t)) = exists_paired_all t | |
| 524 | | exists_paired_all _ = false; | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 525 | val ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 526 | simpset_of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 527 |        (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 528 |         addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}]
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 529 |         addsimprocs [@{simproc unit_eq}]);
 | 
| 11838 | 530 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 531 | fun split_all_tac ctxt = SUBGOAL (fn (t, i) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 532 | if exists_paired_all t then safe_full_simp_tac (put_simpset ss ctxt) i else no_tac); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 533 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 534 | fun unsafe_split_all_tac ctxt = SUBGOAL (fn (t, i) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 535 | if exists_paired_all t then full_simp_tac (put_simpset ss ctxt) i else no_tac); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 536 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 537 | fun split_all ctxt th = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 538 | if exists_paired_all (Thm.prop_of th) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 539 | then full_simplify (put_simpset ss ctxt) th else th; | 
| 11838 | 540 | end; | 
| 26340 | 541 | *} | 
| 11838 | 542 | |
| 51703 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 wenzelm parents: 
51392diff
changeset | 543 | setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac)) *}
 | 
| 11838 | 544 | |
| 47740 
a8989fe9a3a5
added "no_atp"s for extremely prolific, useless facts for ATPs
 blanchet parents: 
46950diff
changeset | 545 | lemma split_paired_All [simp, no_atp]: "(ALL x. P x) = (ALL a b. P (a, b))" | 
| 11838 | 546 |   -- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *}
 | 
| 547 | by fast | |
| 548 | ||
| 47740 
a8989fe9a3a5
added "no_atp"s for extremely prolific, useless facts for ATPs
 blanchet parents: 
46950diff
changeset | 549 | lemma split_paired_Ex [simp, no_atp]: "(EX x. P x) = (EX a b. P (a, b))" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 550 | by fast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 551 | |
| 47740 
a8989fe9a3a5
added "no_atp"s for extremely prolific, useless facts for ATPs
 blanchet parents: 
46950diff
changeset | 552 | lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))" | 
| 11838 | 553 |   -- {* Can't be added to simpset: loops! *}
 | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 554 | by (simp add: split_eta) | 
| 11838 | 555 | |
| 556 | text {*
 | |
| 557 |   Simplification procedure for @{thm [source] cond_split_eta}.  Using
 | |
| 558 |   @{thm [source] split_eta} as a rewrite rule is not general enough,
 | |
| 559 |   and using @{thm [source] cond_split_eta} directly would render some
 | |
| 560 |   existing proofs very inefficient; similarly for @{text
 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 561 | split_beta}. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 562 | *} | 
| 11838 | 563 | |
| 26480 | 564 | ML {*
 | 
| 11838 | 565 | local | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 566 | val cond_split_eta_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 567 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms cond_split_eta});
 | 
| 35364 | 568 | fun Pair_pat k 0 (Bound m) = (m = k) | 
| 569 |     | Pair_pat k i (Const (@{const_name Pair},  _) $ Bound m $ t) =
 | |
| 570 | i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t | |
| 571 | | Pair_pat _ _ _ = false; | |
| 572 | fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t | |
| 573 | | no_args k i (t $ u) = no_args k i t andalso no_args k i u | |
| 574 | | no_args k i (Bound m) = m < k orelse m > k + i | |
| 575 | | no_args _ _ _ = true; | |
| 576 | fun split_pat tp i (Abs (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 577 |     | split_pat tp i (Const (@{const_name case_prod}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t
 | 
| 35364 | 578 | | split_pat tp i _ = NONE; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 579 | fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] [] | 
| 35364 | 580 | (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 581 | (K (simp_tac (put_simpset cond_split_eta_ss ctxt) 1))); | 
| 11838 | 582 | |
| 35364 | 583 | fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t | 
| 584 | | beta_term_pat k i (t $ u) = | |
| 585 | Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u) | |
| 586 | | beta_term_pat k i t = no_args k i t; | |
| 587 | fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg | |
| 588 | | eta_term_pat _ _ _ = false; | |
| 11838 | 589 | fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) | 
| 35364 | 590 | | subst arg k i (t $ u) = | 
| 591 | if Pair_pat k i (t $ u) then incr_boundvars k arg | |
| 592 | else (subst arg k i t $ subst arg k i u) | |
| 593 | | subst arg k i t = t; | |
| 43595 | 594 | in | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 595 |   fun beta_proc ctxt (s as Const (@{const_name case_prod}, _) $ Abs (_, _, t) $ arg) =
 | 
| 11838 | 596 | (case split_pat beta_term_pat 1 t of | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 597 | SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f)) | 
| 15531 | 598 | | NONE => NONE) | 
| 35364 | 599 | | beta_proc _ _ = NONE; | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 600 |   fun eta_proc ctxt (s as Const (@{const_name case_prod}, _) $ Abs (_, _, t)) =
 | 
| 11838 | 601 | (case split_pat eta_term_pat 1 t of | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 602 | SOME (_, ft) => SOME (metaeq ctxt s (let val (f $ arg) = ft in f end)) | 
| 15531 | 603 | | NONE => NONE) | 
| 35364 | 604 | | eta_proc _ _ = NONE; | 
| 11838 | 605 | end; | 
| 606 | *} | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 607 | simproc_setup split_beta ("split f z") = {* fn _ => fn ctxt => fn ct => beta_proc ctxt (term_of ct) *}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 608 | simproc_setup split_eta ("split f") = {* fn _ => fn ctxt => fn ct => eta_proc ctxt (term_of ct) *}
 | 
| 11838 | 609 | |
| 26798 
a9134a089106
split_beta is now declared as monotonicity rule, to allow bounded
 berghofe parents: 
26588diff
changeset | 610 | lemma split_beta [mono]: "(%(x, y). P x y) z = P (fst z) (snd z)" | 
| 11838 | 611 | by (subst surjective_pairing, rule split_conv) | 
| 612 | ||
| 50104 | 613 | lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))" | 
| 614 | by (auto simp: fun_eq_iff) | |
| 615 | ||
| 616 | ||
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 617 | lemma split_split [no_atp]: "R(split c p) = (ALL x y. p = (x, y) --> R(c x y))" | 
| 11838 | 618 |   -- {* For use with @{text split} and the Simplifier. *}
 | 
| 15481 | 619 | by (insert surj_pair [of p], clarify, simp) | 
| 11838 | 620 | |
| 621 | text {*
 | |
| 622 |   @{thm [source] split_split} could be declared as @{text "[split]"}
 | |
| 623 | done after the Splitter has been speeded up significantly; | |
| 624 | precompute the constants involved and don't do anything unless the | |
| 625 | current goal contains one of those constants. | |
| 626 | *} | |
| 627 | ||
| 35828 
46cfc4b8112e
now use "Named_Thms" for "noatp", and renamed "noatp" to "no_atp"
 blanchet parents: 
35427diff
changeset | 628 | lemma split_split_asm [no_atp]: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))" | 
| 14208 | 629 | by (subst split_split, simp) | 
| 11838 | 630 | |
| 631 | text {*
 | |
| 632 |   \medskip @{term split} used as a logical connective or set former.
 | |
| 633 | ||
| 634 |   \medskip These rules are for use with @{text blast}; could instead
 | |
| 40929 
7ff03a5e044f
theorem names generated by the (rep_)datatype command now have mandatory qualifiers
 huffman parents: 
40702diff
changeset | 635 |   call @{text simp} using @{thm [source] prod.split} as rewrite. *}
 | 
| 11838 | 636 | |
| 637 | lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p" | |
| 638 | apply (simp only: split_tupled_all) | |
| 639 | apply (simp (no_asm_simp)) | |
| 640 | done | |
| 641 | ||
| 642 | lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x" | |
| 643 | apply (simp only: split_tupled_all) | |
| 644 | apply (simp (no_asm_simp)) | |
| 645 | done | |
| 646 | ||
| 647 | lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" | |
| 37591 | 648 | by (induct p) auto | 
| 11838 | 649 | |
| 650 | lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" | |
| 37591 | 651 | by (induct p) auto | 
| 11838 | 652 | |
| 653 | lemma splitE2: | |
| 654 | "[| Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R" | |
| 655 | proof - | |
| 656 | assume q: "Q (split P z)" | |
| 657 | assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R" | |
| 658 | show R | |
| 659 | apply (rule r surjective_pairing)+ | |
| 660 | apply (rule split_beta [THEN subst], rule q) | |
| 661 | done | |
| 662 | qed | |
| 663 | ||
| 664 | lemma splitD': "split R (a,b) c ==> R a b c" | |
| 665 | by simp | |
| 666 | ||
| 667 | lemma mem_splitI: "z: c a b ==> z: split c (a, b)" | |
| 668 | by simp | |
| 669 | ||
| 670 | lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p" | |
| 14208 | 671 | by (simp only: split_tupled_all, simp) | 
| 11838 | 672 | |
| 18372 | 673 | lemma mem_splitE: | 
| 37166 | 674 | assumes major: "z \<in> split c p" | 
| 675 | and cases: "\<And>x y. p = (x, y) \<Longrightarrow> z \<in> c x y \<Longrightarrow> Q" | |
| 18372 | 676 | shows Q | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 677 | by (rule major [unfolded case_prod_unfold] cases surjective_pairing)+ | 
| 11838 | 678 | |
| 679 | declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!] | |
| 680 | declare mem_splitE [elim!] splitE' [elim!] splitE [elim!] | |
| 681 | ||
| 26340 | 682 | ML {*
 | 
| 11838 | 683 | local (* filtering with exists_p_split is an essential optimization *) | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 684 |   fun exists_p_split (Const (@{const_name case_prod},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true
 | 
| 11838 | 685 | | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u | 
| 686 | | exists_p_split (Abs (_, _, t)) = exists_p_split t | |
| 687 | | exists_p_split _ = false; | |
| 688 | in | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 689 | fun split_conv_tac ctxt = SUBGOAL (fn (t, i) => | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 690 | if exists_p_split t | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 691 |   then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms split_conv}) i
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 692 | else no_tac); | 
| 11838 | 693 | end; | 
| 26340 | 694 | *} | 
| 695 | ||
| 11838 | 696 | (* This prevents applications of splitE for already splitted arguments leading | 
| 697 | to quite time-consuming computations (in particular for nested tuples) *) | |
| 51703 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 wenzelm parents: 
51392diff
changeset | 698 | setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac)) *}
 | 
| 11838 | 699 | |
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
52435diff
changeset | 700 | lemma split_eta_SetCompr [simp, no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" | 
| 18372 | 701 | by (rule ext) fast | 
| 11838 | 702 | |
| 54147 
97a8ff4e4ac9
killed most "no_atp", to make Sledgehammer more complete
 blanchet parents: 
52435diff
changeset | 703 | lemma split_eta_SetCompr2 [simp, no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P" | 
| 18372 | 704 | by (rule ext) fast | 
| 11838 | 705 | |
| 706 | lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)" | |
| 707 |   -- {* Allows simplifications of nested splits in case of independent predicates. *}
 | |
| 18372 | 708 | by (rule ext) blast | 
| 11838 | 709 | |
| 14337 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 710 | (* Do NOT make this a simp rule as it | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 711 | a) only helps in special situations | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 712 | b) can lead to nontermination in the presence of split_def | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 713 | *) | 
| 
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
 nipkow parents: 
14208diff
changeset | 714 | lemma split_comp_eq: | 
| 20415 | 715 | fixes f :: "'a => 'b => 'c" and g :: "'d => 'a" | 
| 716 | shows "(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))" | |
| 18372 | 717 | by (rule ext) auto | 
| 14101 | 718 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 719 | lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 720 | apply (rule_tac x = "(a, b)" in image_eqI) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 721 | apply auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 722 | done | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 723 | |
| 11838 | 724 | lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)" | 
| 725 | by blast | |
| 726 | ||
| 727 | (* | |
| 728 | the following would be slightly more general, | |
| 729 | but cannot be used as rewrite rule: | |
| 730 | ### Cannot add premise as rewrite rule because it contains (type) unknowns: | |
| 731 | ### ?y = .x | |
| 732 | Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" | |
| 14208 | 733 | by (rtac some_equality 1) | 
| 734 | by ( Simp_tac 1) | |
| 735 | by (split_all_tac 1) | |
| 736 | by (Asm_full_simp_tac 1) | |
| 11838 | 737 | qed "The_split_eq"; | 
| 738 | *) | |
| 739 | ||
| 740 | text {*
 | |
| 741 |   Setup of internal @{text split_rule}.
 | |
| 742 | *} | |
| 743 | ||
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55469diff
changeset | 744 | lemmas case_prodI = prod.case [THEN iffD2] | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 745 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 746 | lemma case_prodI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> case_prod c p" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 747 | by (fact splitI2) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 748 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 749 | lemma case_prodI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> case_prod c p x" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 750 | by (fact splitI2') | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 751 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 752 | lemma case_prodE: "case_prod c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 753 | by (fact splitE) | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 754 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 755 | lemma case_prodE': "case_prod c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 756 | by (fact splitE') | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 757 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 758 | declare case_prodI [intro!] | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 759 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 760 | lemma case_prod_beta: | 
| 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 761 | "case_prod f p = f (fst p) (snd p)" | 
| 37591 | 762 | by (fact split_beta) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
25885diff
changeset | 763 | |
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
55414diff
changeset | 764 | lemma prod_cases3 [cases type]: | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 765 | obtains (fields) a b c where "y = (a, b, c)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 766 | by (cases y, case_tac b) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 767 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 768 | lemma prod_induct3 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 769 | "(!!a b c. P (a, b, c)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 770 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 771 | |
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
55414diff
changeset | 772 | lemma prod_cases4 [cases type]: | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 773 | obtains (fields) a b c d where "y = (a, b, c, d)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 774 | by (cases y, case_tac c) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 775 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 776 | lemma prod_induct4 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 777 | "(!!a b c d. P (a, b, c, d)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 778 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 779 | |
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
55414diff
changeset | 780 | lemma prod_cases5 [cases type]: | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 781 | obtains (fields) a b c d e where "y = (a, b, c, d, e)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 782 | by (cases y, case_tac d) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 783 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 784 | lemma prod_induct5 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 785 | "(!!a b c d e. P (a, b, c, d, e)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 786 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 787 | |
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
55414diff
changeset | 788 | lemma prod_cases6 [cases type]: | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 789 | obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 790 | by (cases y, case_tac e) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 791 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 792 | lemma prod_induct6 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 793 | "(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 794 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 795 | |
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
55414diff
changeset | 796 | lemma prod_cases7 [cases type]: | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 797 | obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 798 | by (cases y, case_tac f) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 799 | |
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 800 | lemma prod_induct7 [case_names fields, induct type]: | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 801 | "(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x" | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 802 | by (cases x) blast | 
| 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 803 | |
| 37166 | 804 | lemma split_def: | 
| 805 | "split = (\<lambda>c p. c (fst p) (snd p))" | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 806 | by (fact case_prod_unfold) | 
| 37166 | 807 | |
| 808 | definition internal_split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
 | |
| 809 | "internal_split == split" | |
| 810 | ||
| 811 | lemma internal_split_conv: "internal_split c (a, b) = c a b" | |
| 812 | by (simp only: internal_split_def split_conv) | |
| 813 | ||
| 48891 | 814 | ML_file "Tools/split_rule.ML" | 
| 37166 | 815 | setup Split_Rule.setup | 
| 816 | ||
| 817 | hide_const internal_split | |
| 818 | ||
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 819 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 820 | subsubsection {* Derived operations *}
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 821 | |
| 37387 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
37278diff
changeset | 822 | definition curry    :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c" where
 | 
| 
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
 haftmann parents: 
37278diff
changeset | 823 | "curry = (\<lambda>c x y. c (x, y))" | 
| 37166 | 824 | |
| 825 | lemma curry_conv [simp, code]: "curry f a b = f (a, b)" | |
| 826 | by (simp add: curry_def) | |
| 827 | ||
| 828 | lemma curryI [intro!]: "f (a, b) \<Longrightarrow> curry f a b" | |
| 829 | by (simp add: curry_def) | |
| 830 | ||
| 831 | lemma curryD [dest!]: "curry f a b \<Longrightarrow> f (a, b)" | |
| 832 | by (simp add: curry_def) | |
| 833 | ||
| 834 | lemma curryE: "curry f a b \<Longrightarrow> (f (a, b) \<Longrightarrow> Q) \<Longrightarrow> Q" | |
| 835 | by (simp add: curry_def) | |
| 836 | ||
| 837 | lemma curry_split [simp]: "curry (split f) = f" | |
| 838 | by (simp add: curry_def split_def) | |
| 839 | ||
| 840 | lemma split_curry [simp]: "split (curry f) = f" | |
| 841 | by (simp add: curry_def split_def) | |
| 842 | ||
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
54147diff
changeset | 843 | lemma curry_K: "curry (\<lambda>x. c) = (\<lambda>x y. c)" | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
54147diff
changeset | 844 | by(simp add: fun_eq_iff) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
54147diff
changeset | 845 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 846 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 847 | The composition-uncurry combinator. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 848 | *} | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 849 | |
| 37751 | 850 | notation fcomp (infixl "\<circ>>" 60) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 851 | |
| 37751 | 852 | definition scomp :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd" (infixl "\<circ>\<rightarrow>" 60) where
 | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 853 | "f \<circ>\<rightarrow> g = (\<lambda>x. case_prod g (f x))" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 854 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 855 | lemma scomp_unfold: "scomp = (\<lambda>f g x. g (fst (f x)) (snd (f x)))" | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 856 | by (simp add: fun_eq_iff scomp_def case_prod_unfold) | 
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37591diff
changeset | 857 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 858 | lemma scomp_apply [simp]: "(f \<circ>\<rightarrow> g) x = case_prod g (f x)" | 
| 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 859 | by (simp add: scomp_unfold case_prod_unfold) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 860 | |
| 37751 | 861 | lemma Pair_scomp: "Pair x \<circ>\<rightarrow> f = f x" | 
| 44921 | 862 | by (simp add: fun_eq_iff) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 863 | |
| 37751 | 864 | lemma scomp_Pair: "x \<circ>\<rightarrow> Pair = x" | 
| 44921 | 865 | by (simp add: fun_eq_iff) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 866 | |
| 37751 | 867 | lemma scomp_scomp: "(f \<circ>\<rightarrow> g) \<circ>\<rightarrow> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>\<rightarrow> h)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 868 | by (simp add: fun_eq_iff scomp_unfold) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 869 | |
| 37751 | 870 | lemma scomp_fcomp: "(f \<circ>\<rightarrow> g) \<circ>> h = f \<circ>\<rightarrow> (\<lambda>x. g x \<circ>> h)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 871 | by (simp add: fun_eq_iff scomp_unfold fcomp_def) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 872 | |
| 37751 | 873 | lemma fcomp_scomp: "(f \<circ>> g) \<circ>\<rightarrow> h = f \<circ>> (g \<circ>\<rightarrow> h)" | 
| 44921 | 874 | by (simp add: fun_eq_iff scomp_unfold) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 875 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 876 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52143diff
changeset | 877 | constant scomp \<rightharpoonup> (Eval) infixl 3 "#->" | 
| 31202 
52d332f8f909
pretty printing of functional combinators for evaluation code
 haftmann parents: 
30924diff
changeset | 878 | |
| 37751 | 879 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 880 | no_notation scomp (infixl "\<circ>\<rightarrow>" 60) | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 881 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 882 | text {*
 | 
| 55932 | 883 |   @{term map_prod} --- action of the product functor upon
 | 
| 36664 
6302f9ad7047
repaired comments where SOMEthing went utterly wrong (cf. 2b04504fcb69)
 krauss parents: 
36622diff
changeset | 884 | functions. | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 885 | *} | 
| 21195 | 886 | |
| 55932 | 887 | definition map_prod :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'd" where
 | 
| 888 | "map_prod f g = (\<lambda>(x, y). (f x, g y))" | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 889 | |
| 55932 | 890 | lemma map_prod_simp [simp, code]: | 
| 891 | "map_prod f g (a, b) = (f a, g b)" | |
| 892 | by (simp add: map_prod_def) | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 893 | |
| 55932 | 894 | functor map_prod: map_prod | 
| 44921 | 895 | by (auto simp add: split_paired_all) | 
| 37278 | 896 | |
| 55932 | 897 | lemma fst_map_prod [simp]: | 
| 898 | "fst (map_prod f g x) = f (fst x)" | |
| 40607 | 899 | by (cases x) simp_all | 
| 37278 | 900 | |
| 40607 | 901 | lemma snd_prod_fun [simp]: | 
| 55932 | 902 | "snd (map_prod f g x) = g (snd x)" | 
| 40607 | 903 | by (cases x) simp_all | 
| 37278 | 904 | |
| 55932 | 905 | lemma fst_comp_map_prod [simp]: | 
| 906 | "fst \<circ> map_prod f g = f \<circ> fst" | |
| 40607 | 907 | by (rule ext) simp_all | 
| 37278 | 908 | |
| 55932 | 909 | lemma snd_comp_map_prod [simp]: | 
| 910 | "snd \<circ> map_prod f g = g \<circ> snd" | |
| 40607 | 911 | by (rule ext) simp_all | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 912 | |
| 55932 | 913 | lemma map_prod_compose: | 
| 914 | "map_prod (f1 o f2) (g1 o g2) = (map_prod f1 g1 o map_prod f2 g2)" | |
| 915 | by (rule ext) (simp add: map_prod.compositionality comp_def) | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 916 | |
| 55932 | 917 | lemma map_prod_ident [simp]: | 
| 918 | "map_prod (%x. x) (%y. y) = (%z. z)" | |
| 919 | by (rule ext) (simp add: map_prod.identity) | |
| 40607 | 920 | |
| 55932 | 921 | lemma map_prod_imageI [intro]: | 
| 922 | "(a, b) \<in> R \<Longrightarrow> (f a, g b) \<in> map_prod f g ` R" | |
| 40607 | 923 | by (rule image_eqI) simp_all | 
| 21195 | 924 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 925 | lemma prod_fun_imageE [elim!]: | 
| 55932 | 926 | assumes major: "c \<in> map_prod f g ` R" | 
| 40607 | 927 | and cases: "\<And>x y. c = (f x, g y) \<Longrightarrow> (x, y) \<in> R \<Longrightarrow> P" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 928 | shows P | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 929 | apply (rule major [THEN imageE]) | 
| 37166 | 930 | apply (case_tac x) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 931 | apply (rule cases) | 
| 40607 | 932 | apply simp_all | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 933 | done | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 934 | |
| 37166 | 935 | definition apfst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b" where
 | 
| 55932 | 936 | "apfst f = map_prod f id" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 937 | |
| 37166 | 938 | definition apsnd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c" where
 | 
| 55932 | 939 | "apsnd f = map_prod id f" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 940 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 941 | lemma apfst_conv [simp, code]: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 942 | "apfst f (x, y) = (f x, y)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 943 | by (simp add: apfst_def) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 944 | |
| 33638 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 945 | lemma apsnd_conv [simp, code]: | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 946 | "apsnd f (x, y) = (x, f y)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 947 | by (simp add: apsnd_def) | 
| 21195 | 948 | |
| 33594 | 949 | lemma fst_apfst [simp]: | 
| 950 | "fst (apfst f x) = f (fst x)" | |
| 951 | by (cases x) simp | |
| 952 | ||
| 51173 | 953 | lemma fst_comp_apfst [simp]: | 
| 954 | "fst \<circ> apfst f = f \<circ> fst" | |
| 955 | by (simp add: fun_eq_iff) | |
| 956 | ||
| 33594 | 957 | lemma fst_apsnd [simp]: | 
| 958 | "fst (apsnd f x) = fst x" | |
| 959 | by (cases x) simp | |
| 960 | ||
| 51173 | 961 | lemma fst_comp_apsnd [simp]: | 
| 962 | "fst \<circ> apsnd f = fst" | |
| 963 | by (simp add: fun_eq_iff) | |
| 964 | ||
| 33594 | 965 | lemma snd_apfst [simp]: | 
| 966 | "snd (apfst f x) = snd x" | |
| 967 | by (cases x) simp | |
| 968 | ||
| 51173 | 969 | lemma snd_comp_apfst [simp]: | 
| 970 | "snd \<circ> apfst f = snd" | |
| 971 | by (simp add: fun_eq_iff) | |
| 972 | ||
| 33594 | 973 | lemma snd_apsnd [simp]: | 
| 974 | "snd (apsnd f x) = f (snd x)" | |
| 975 | by (cases x) simp | |
| 976 | ||
| 51173 | 977 | lemma snd_comp_apsnd [simp]: | 
| 978 | "snd \<circ> apsnd f = f \<circ> snd" | |
| 979 | by (simp add: fun_eq_iff) | |
| 980 | ||
| 33594 | 981 | lemma apfst_compose: | 
| 982 | "apfst f (apfst g x) = apfst (f \<circ> g) x" | |
| 983 | by (cases x) simp | |
| 984 | ||
| 985 | lemma apsnd_compose: | |
| 986 | "apsnd f (apsnd g x) = apsnd (f \<circ> g) x" | |
| 987 | by (cases x) simp | |
| 988 | ||
| 989 | lemma apfst_apsnd [simp]: | |
| 990 | "apfst f (apsnd g x) = (f (fst x), g (snd x))" | |
| 991 | by (cases x) simp | |
| 992 | ||
| 993 | lemma apsnd_apfst [simp]: | |
| 994 | "apsnd f (apfst g x) = (g (fst x), f (snd x))" | |
| 995 | by (cases x) simp | |
| 996 | ||
| 997 | lemma apfst_id [simp] : | |
| 998 | "apfst id = id" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 999 | by (simp add: fun_eq_iff) | 
| 33594 | 1000 | |
| 1001 | lemma apsnd_id [simp] : | |
| 1002 | "apsnd id = id" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39272diff
changeset | 1003 | by (simp add: fun_eq_iff) | 
| 33594 | 1004 | |
| 1005 | lemma apfst_eq_conv [simp]: | |
| 1006 | "apfst f x = apfst g x \<longleftrightarrow> f (fst x) = g (fst x)" | |
| 1007 | by (cases x) simp | |
| 1008 | ||
| 1009 | lemma apsnd_eq_conv [simp]: | |
| 1010 | "apsnd f x = apsnd g x \<longleftrightarrow> f (snd x) = g (snd x)" | |
| 1011 | by (cases x) simp | |
| 1012 | ||
| 33638 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 1013 | lemma apsnd_apfst_commute: | 
| 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 1014 | "apsnd f (apfst g p) = apfst g (apsnd f p)" | 
| 
548a34929e98
Renamed upd_snd_conv to apsnd_conv to be consistent with apfst_conv; Added apsnd_apfst_commute
 hoelzl parents: 
33594diff
changeset | 1015 | by simp | 
| 21195 | 1016 | |
| 56626 | 1017 | context | 
| 1018 | begin | |
| 1019 | ||
| 1020 | local_setup {* Local_Theory.map_naming (Name_Space.mandatory_path "prod") *}
 | |
| 1021 | ||
| 56545 | 1022 | definition swap :: "'a \<times> 'b \<Rightarrow> 'b \<times> 'a" | 
| 1023 | where | |
| 1024 | "swap p = (snd p, fst p)" | |
| 1025 | ||
| 56626 | 1026 | end | 
| 1027 | ||
| 56545 | 1028 | lemma swap_simp [simp]: | 
| 56626 | 1029 | "prod.swap (x, y) = (y, x)" | 
| 1030 | by (simp add: prod.swap_def) | |
| 56545 | 1031 | |
| 1032 | lemma pair_in_swap_image [simp]: | |
| 56626 | 1033 | "(y, x) \<in> prod.swap ` A \<longleftrightarrow> (x, y) \<in> A" | 
| 56545 | 1034 | by (auto intro!: image_eqI) | 
| 1035 | ||
| 1036 | lemma inj_swap [simp]: | |
| 56626 | 1037 | "inj_on prod.swap A" | 
| 1038 | by (rule inj_onI) auto | |
| 1039 | ||
| 1040 | lemma swap_inj_on: | |
| 1041 | "inj_on (\<lambda>(i, j). (j, i)) A" | |
| 1042 | by (rule inj_onI) auto | |
| 56545 | 1043 | |
| 1044 | lemma case_swap [simp]: | |
| 56626 | 1045 | "(case prod.swap p of (y, x) \<Rightarrow> f x y) = (case p of (x, y) \<Rightarrow> f x y)" | 
| 56545 | 1046 | by (cases p) simp | 
| 1047 | ||
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1048 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1049 | Disjoint union of a family of sets -- Sigma. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1050 | *} | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1051 | |
| 45986 
c9e50153e5ae
moved various set operations to theory Set (resp. Product_Type)
 haftmann parents: 
45696diff
changeset | 1052 | definition Sigma :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<times> 'b) set" where
 | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1053 |   Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1054 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1055 | abbreviation | 
| 45986 
c9e50153e5ae
moved various set operations to theory Set (resp. Product_Type)
 haftmann parents: 
45696diff
changeset | 1056 |   Times :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set"
 | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1057 | (infixr "<*>" 80) where | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1058 | "A <*> B == Sigma A (%_. B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1059 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1060 | notation (xsymbols) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1061 | Times (infixr "\<times>" 80) | 
| 15394 | 1062 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1063 | notation (HTML output) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1064 | Times (infixr "\<times>" 80) | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1065 | |
| 45662 
4f7c05990420
Hide Product_Type.Times - too precious an identifier
 nipkow parents: 
45607diff
changeset | 1066 | hide_const (open) Times | 
| 
4f7c05990420
Hide Product_Type.Times - too precious an identifier
 nipkow parents: 
45607diff
changeset | 1067 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1068 | syntax | 
| 35115 | 1069 |   "_Sigma" :: "[pttrn, 'a set, 'b set] => ('a * 'b) set"  ("(3SIGMA _:_./ _)" [0, 0, 10] 10)
 | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1070 | translations | 
| 35115 | 1071 | "SIGMA x:A. B" == "CONST Sigma A (%x. B)" | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1072 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1073 | lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1074 | by (unfold Sigma_def) blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1075 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1076 | lemma SigmaE [elim!]: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1077 | "[| c: Sigma A B; | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1078 | !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1079 | |] ==> P" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1080 |   -- {* The general elimination rule. *}
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1081 | by (unfold Sigma_def) blast | 
| 20588 | 1082 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1083 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1084 |   Elimination of @{term "(a, b) : A \<times> B"} -- introduces no
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1085 | eigenvariables. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1086 | *} | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1087 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1088 | lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1089 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1090 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1091 | lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1092 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1093 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1094 | lemma SigmaE2: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1095 | "[| (a, b) : Sigma A B; | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1096 | [| a:A; b:B(a) |] ==> P | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1097 | |] ==> P" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1098 | by blast | 
| 20588 | 1099 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1100 | lemma Sigma_cong: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1101 | "\<lbrakk>A = B; !!x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk> | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1102 | \<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1103 | by auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1104 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1105 | lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1106 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1107 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1108 | lemma Sigma_empty1 [simp]: "Sigma {} B = {}"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1109 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1110 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1111 | lemma Sigma_empty2 [simp]: "A <*> {} = {}"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1112 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1113 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1114 | lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1115 | by auto | 
| 21908 | 1116 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1117 | lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1118 | by auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1119 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1120 | lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1121 | by auto | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1122 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1123 | lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1124 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1125 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1126 | lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1127 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1128 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1129 | lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1130 | by (blast elim: equalityE) | 
| 20588 | 1131 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1132 | lemma SetCompr_Sigma_eq: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1133 | "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1134 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1135 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1136 | lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q"
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1137 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1138 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1139 | lemma UN_Times_distrib: | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1140 | "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1141 |   -- {* Suggested by Pierre Chartier *}
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1142 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1143 | |
| 47740 
a8989fe9a3a5
added "no_atp"s for extremely prolific, useless facts for ATPs
 blanchet parents: 
46950diff
changeset | 1144 | lemma split_paired_Ball_Sigma [simp, no_atp]: | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1145 | "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1146 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1147 | |
| 47740 
a8989fe9a3a5
added "no_atp"s for extremely prolific, useless facts for ATPs
 blanchet parents: 
46950diff
changeset | 1148 | lemma split_paired_Bex_Sigma [simp, no_atp]: | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1149 | "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1150 | by blast | 
| 21908 | 1151 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1152 | lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1153 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1154 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1155 | lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1156 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1157 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1158 | lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1159 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1160 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1161 | lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1162 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1163 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1164 | lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1165 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1166 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1167 | lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1168 | by blast | 
| 21908 | 1169 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1170 | lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)" | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1171 | by blast | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1172 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1173 | text {*
 | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1174 | Non-dependent versions are needed to avoid the need for higher-order | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1175 | matching, especially when the rules are re-oriented. | 
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1176 | *} | 
| 21908 | 1177 | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1178 | lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)" | 
| 56545 | 1179 | by (fact Sigma_Un_distrib1) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1180 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1181 | lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)" | 
| 56545 | 1182 | by (fact Sigma_Int_distrib1) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1183 | |
| 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1184 | lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)" | 
| 56545 | 1185 | by (fact Sigma_Diff_distrib1) | 
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1186 | |
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1187 | lemma Times_empty[simp]: "A \<times> B = {} \<longleftrightarrow> A = {} \<or> B = {}"
 | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1188 | by auto | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1189 | |
| 50104 | 1190 | lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))"
 | 
| 1191 | by auto | |
| 1192 | ||
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1193 | lemma fst_image_times[simp]: "fst ` (A \<times> B) = (if B = {} then {} else A)"
 | 
| 44921 | 1194 | by force | 
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1195 | |
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1196 | lemma snd_image_times[simp]: "snd ` (A \<times> B) = (if A = {} then {} else B)"
 | 
| 44921 | 1197 | by force | 
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1198 | |
| 56545 | 1199 | lemma vimage_fst: | 
| 1200 | "fst -` A = A \<times> UNIV" | |
| 1201 | by auto | |
| 1202 | ||
| 1203 | lemma vimage_snd: | |
| 1204 | "snd -` A = UNIV \<times> A" | |
| 1205 | by auto | |
| 1206 | ||
| 28719 | 1207 | lemma insert_times_insert[simp]: | 
| 1208 | "insert a A \<times> insert b B = | |
| 1209 | insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)" | |
| 1210 | by blast | |
| 26358 
d6a508c16908
Product_Type.apfst and Product_Type.apsnd; mbind combinator; tuned
 haftmann parents: 
26340diff
changeset | 1211 | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: 
33089diff
changeset | 1212 | lemma vimage_Times: "f -` (A \<times> B) = ((fst \<circ> f) -` A) \<inter> ((snd \<circ> f) -` B)" | 
| 47988 | 1213 | apply auto | 
| 1214 | apply (case_tac "f x") | |
| 1215 | apply auto | |
| 1216 | done | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: 
33089diff
changeset | 1217 | |
| 50104 | 1218 | lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)" | 
| 1219 | by auto | |
| 1220 | ||
| 56626 | 1221 | lemma product_swap: | 
| 1222 | "prod.swap ` (A \<times> B) = B \<times> A" | |
| 1223 | by (auto simp add: set_eq_iff) | |
| 35822 | 1224 | |
| 1225 | lemma swap_product: | |
| 56626 | 1226 | "(\<lambda>(i, j). (j, i)) ` (A \<times> B) = B \<times> A" | 
| 1227 | by (auto simp add: set_eq_iff) | |
| 35822 | 1228 | |
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1229 | lemma image_split_eq_Sigma: | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1230 |   "(\<lambda>x. (f x, g x)) ` A = Sigma (f ` A) (\<lambda>x. g ` (f -` {x} \<inter> A))"
 | 
| 46128 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1231 | proof (safe intro!: imageI) | 
| 36622 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1232 | fix a b assume *: "a \<in> A" "b \<in> A" and eq: "f a = f b" | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1233 | show "(f b, g a) \<in> (\<lambda>x. (f x, g x)) ` A" | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1234 | using * eq[symmetric] by auto | 
| 
e393a91f86df
Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
 hoelzl parents: 
36176diff
changeset | 1235 | qed simp_all | 
| 35822 | 1236 | |
| 46128 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1237 | definition product :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
 | 
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1238 | [code_abbrev]: "product A B = A \<times> B" | 
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1239 | |
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1240 | hide_const (open) product | 
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1241 | |
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1242 | lemma member_product: | 
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1243 | "x \<in> Product_Type.product A B \<longleftrightarrow> x \<in> A \<times> B" | 
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1244 | by (simp add: product_def) | 
| 
53e7cc599f58
interaction of set operations for execution and membership predicate
 haftmann parents: 
46028diff
changeset | 1245 | |
| 55932 | 1246 | text {* The following @{const map_prod} lemmas are due to Joachim Breitner: *}
 | 
| 40607 | 1247 | |
| 55932 | 1248 | lemma map_prod_inj_on: | 
| 40607 | 1249 | assumes "inj_on f A" and "inj_on g B" | 
| 55932 | 1250 | shows "inj_on (map_prod f g) (A \<times> B)" | 
| 40607 | 1251 | proof (rule inj_onI) | 
| 1252 | fix x :: "'a \<times> 'c" and y :: "'a \<times> 'c" | |
| 1253 | assume "x \<in> A \<times> B" hence "fst x \<in> A" and "snd x \<in> B" by auto | |
| 1254 | assume "y \<in> A \<times> B" hence "fst y \<in> A" and "snd y \<in> B" by auto | |
| 55932 | 1255 | assume "map_prod f g x = map_prod f g y" | 
| 1256 | hence "fst (map_prod f g x) = fst (map_prod f g y)" by (auto) | |
| 40607 | 1257 | hence "f (fst x) = f (fst y)" by (cases x,cases y,auto) | 
| 1258 | with `inj_on f A` and `fst x \<in> A` and `fst y \<in> A` | |
| 1259 | have "fst x = fst y" by (auto dest:dest:inj_onD) | |
| 55932 | 1260 | moreover from `map_prod f g x = map_prod f g y` | 
| 1261 | have "snd (map_prod f g x) = snd (map_prod f g y)" by (auto) | |
| 40607 | 1262 | hence "g (snd x) = g (snd y)" by (cases x,cases y,auto) | 
| 1263 | with `inj_on g B` and `snd x \<in> B` and `snd y \<in> B` | |
| 1264 | have "snd x = snd y" by (auto dest:dest:inj_onD) | |
| 1265 | ultimately show "x = y" by(rule prod_eqI) | |
| 1266 | qed | |
| 1267 | ||
| 55932 | 1268 | lemma map_prod_surj: | 
| 40702 | 1269 | fixes f :: "'a \<Rightarrow> 'b" and g :: "'c \<Rightarrow> 'd" | 
| 40607 | 1270 | assumes "surj f" and "surj g" | 
| 55932 | 1271 | shows "surj (map_prod f g)" | 
| 40607 | 1272 | unfolding surj_def | 
| 1273 | proof | |
| 1274 | fix y :: "'b \<times> 'd" | |
| 1275 | from `surj f` obtain a where "fst y = f a" by (auto elim:surjE) | |
| 1276 | moreover | |
| 1277 | from `surj g` obtain b where "snd y = g b" by (auto elim:surjE) | |
| 55932 | 1278 | ultimately have "(fst y, snd y) = map_prod f g (a,b)" by auto | 
| 1279 | thus "\<exists>x. y = map_prod f g x" by auto | |
| 40607 | 1280 | qed | 
| 1281 | ||
| 55932 | 1282 | lemma map_prod_surj_on: | 
| 40607 | 1283 | assumes "f ` A = A'" and "g ` B = B'" | 
| 55932 | 1284 | shows "map_prod f g ` (A \<times> B) = A' \<times> B'" | 
| 40607 | 1285 | unfolding image_def | 
| 1286 | proof(rule set_eqI,rule iffI) | |
| 1287 | fix x :: "'a \<times> 'c" | |
| 55932 | 1288 |   assume "x \<in> {y\<Colon>'a \<times> 'c. \<exists>x\<Colon>'b \<times> 'd\<in>A \<times> B. y = map_prod f g x}"
 | 
| 1289 | then obtain y where "y \<in> A \<times> B" and "x = map_prod f g y" by blast | |
| 40607 | 1290 | from `image f A = A'` and `y \<in> A \<times> B` have "f (fst y) \<in> A'" by auto | 
| 1291 | moreover from `image g B = B'` and `y \<in> A \<times> B` have "g (snd y) \<in> B'" by auto | |
| 1292 | ultimately have "(f (fst y), g (snd y)) \<in> (A' \<times> B')" by auto | |
| 55932 | 1293 | with `x = map_prod f g y` show "x \<in> A' \<times> B'" by (cases y, auto) | 
| 40607 | 1294 | next | 
| 1295 | fix x :: "'a \<times> 'c" | |
| 1296 | assume "x \<in> A' \<times> B'" hence "fst x \<in> A'" and "snd x \<in> B'" by auto | |
| 1297 | from `image f A = A'` and `fst x \<in> A'` have "fst x \<in> image f A" by auto | |
| 1298 | then obtain a where "a \<in> A" and "fst x = f a" by (rule imageE) | |
| 1299 | moreover from `image g B = B'` and `snd x \<in> B'` | |
| 1300 | obtain b where "b \<in> B" and "snd x = g b" by auto | |
| 55932 | 1301 | ultimately have "(fst x, snd x) = map_prod f g (a,b)" by auto | 
| 40607 | 1302 | moreover from `a \<in> A` and `b \<in> B` have "(a , b) \<in> A \<times> B" by auto | 
| 55932 | 1303 | ultimately have "\<exists>y \<in> A \<times> B. x = map_prod f g y" by auto | 
| 1304 |   thus "x \<in> {x. \<exists>y \<in> A \<times> B. x = map_prod f g y}" by auto
 | |
| 40607 | 1305 | qed | 
| 1306 | ||
| 21908 | 1307 | |
| 49764 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1308 | subsection {* Simproc for rewriting a set comprehension into a pointfree expression *}
 | 
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1309 | |
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1310 | ML_file "Tools/set_comprehension_pointfree.ML" | 
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1311 | |
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1312 | setup {*
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51703diff
changeset | 1313 | Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs | 
| 49764 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1314 |     [Raw_Simplifier.make_simproc {name = "set comprehension", lhss = [@{cpat "Collect ?P"}],
 | 
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1315 | proc = K Set_Comprehension_Pointfree.code_simproc, identifier = []}]) | 
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1316 | *} | 
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1317 | |
| 
9979d64b8016
moving simproc from Finite_Set to more appropriate Product_Type theory
 bulwahn parents: 
48891diff
changeset | 1318 | |
| 37166 | 1319 | subsection {* Inductively defined sets *}
 | 
| 15394 | 1320 | |
| 56512 | 1321 | (* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *)
 | 
| 1322 | simproc_setup Collect_mem ("Collect t") = {*
 | |
| 1323 | fn _ => fn ctxt => fn ct => | |
| 1324 | (case term_of ct of | |
| 1325 |       S as Const (@{const_name Collect}, Type (@{type_name fun}, [_, T])) $ t =>
 | |
| 1326 | let val (u, _, ps) = HOLogic.strip_psplits t in | |
| 1327 | (case u of | |
| 1328 |             (c as Const (@{const_name Set.member}, _)) $ q $ S' =>
 | |
| 1329 | (case try (HOLogic.strip_ptuple ps) q of | |
| 1330 | NONE => NONE | |
| 1331 | | SOME ts => | |
| 1332 | if not (Term.is_open S') andalso | |
| 1333 | ts = map Bound (length ps downto 0) | |
| 1334 | then | |
| 1335 | let val simp = | |
| 1336 | full_simp_tac (put_simpset HOL_basic_ss ctxt | |
| 1337 |                         addsimps [@{thm split_paired_all}, @{thm split_conv}]) 1
 | |
| 1338 | in | |
| 1339 | SOME (Goal.prove ctxt [] [] | |
| 1340 |                         (Const (@{const_name Pure.eq}, T --> T --> propT) $ S $ S')
 | |
| 1341 | (K (EVERY | |
| 1342 |                           [rtac eq_reflection 1, rtac @{thm subset_antisym} 1,
 | |
| 1343 | rtac subsetI 1, dtac CollectD 1, simp, | |
| 1344 | rtac subsetI 1, rtac CollectI 1, simp]))) | |
| 1345 | end | |
| 1346 | else NONE) | |
| 1347 | | _ => NONE) | |
| 1348 | end | |
| 1349 | | _ => NONE) | |
| 1350 | *} | |
| 48891 | 1351 | ML_file "Tools/inductive_set.ML" | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24286diff
changeset | 1352 | |
| 37166 | 1353 | |
| 1354 | subsection {* Legacy theorem bindings and duplicates *}
 | |
| 1355 | ||
| 1356 | lemma PairE: | |
| 1357 | obtains x y where "p = (x, y)" | |
| 1358 | by (fact prod.exhaust) | |
| 1359 | ||
| 1360 | lemmas Pair_eq = prod.inject | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 1361 | lemmas fst_conv = prod.sel(1) | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 1362 | lemmas snd_conv = prod.sel(2) | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 1363 | lemmas pair_collapse = prod.collapse | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
54630diff
changeset | 1364 | lemmas split = split_conv | 
| 44066 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
43866diff
changeset | 1365 | lemmas Pair_fst_snd_eq = prod_eq_iff | 
| 
d74182c93f04
rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
 huffman parents: 
43866diff
changeset | 1366 | |
| 45204 
5e4a1270c000
hide typedef-generated constants Product_Type.prod and Sum_Type.sum
 huffman parents: 
44921diff
changeset | 1367 | hide_const (open) prod | 
| 
5e4a1270c000
hide typedef-generated constants Product_Type.prod and Sum_Type.sum
 huffman parents: 
44921diff
changeset | 1368 | |
| 10213 | 1369 | end |