| author | chaieb | 
| Wed, 22 Aug 2007 17:13:41 +0200 | |
| changeset 24403 | b7c3ee2ca184 | 
| parent 24349 | 0dd8782fb02d | 
| child 24625 | 0398a5e802d3 | 
| permissions | -rw-r--r-- | 
| 7700 | 1  | 
(* Title: HOL/Inductive.thy  | 
2  | 
ID: $Id$  | 
|
| 10402 | 3  | 
Author: Markus Wenzel, TU Muenchen  | 
| 11688 | 4  | 
*)  | 
| 10727 | 5  | 
|
| 11688 | 6  | 
header {* Support for inductive sets and types *}
 | 
| 1187 | 7  | 
|
| 15131 | 8  | 
theory Inductive  | 
| 23708 | 9  | 
imports FixedPoint Product_Type Sum_Type  | 
| 16417 | 10  | 
uses  | 
| 10402 | 11  | 
  ("Tools/inductive_package.ML")
 | 
| 23734 | 12  | 
  ("Tools/inductive_set_package.ML")
 | 
| 13705 | 13  | 
  ("Tools/inductive_realizer.ML")
 | 
| 
12437
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
14  | 
  ("Tools/inductive_codegen.ML")
 | 
| 10402 | 15  | 
  ("Tools/datatype_aux.ML")
 | 
16  | 
  ("Tools/datatype_prop.ML")
 | 
|
17  | 
  ("Tools/datatype_rep_proofs.ML")
 | 
|
18  | 
  ("Tools/datatype_abs_proofs.ML")
 | 
|
| 13469 | 19  | 
  ("Tools/datatype_realizer.ML")
 | 
| 19599 | 20  | 
  ("Tools/datatype_hooks.ML")
 | 
| 22783 | 21  | 
  ("Tools/datatype_case.ML")
 | 
| 10402 | 22  | 
  ("Tools/datatype_package.ML")
 | 
| 
12437
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
23  | 
  ("Tools/datatype_codegen.ML")
 | 
| 15131 | 24  | 
  ("Tools/primrec_package.ML")
 | 
25  | 
begin  | 
|
| 10727 | 26  | 
|
| 23734 | 27  | 
subsection {* Inductive predicates and sets *}
 | 
| 11688 | 28  | 
|
29  | 
text {* Inversion of injective functions. *}
 | 
|
| 11436 | 30  | 
|
31  | 
constdefs  | 
|
32  | 
  myinv :: "('a => 'b) => ('b => 'a)"
 | 
|
33  | 
"myinv (f :: 'a => 'b) == \<lambda>y. THE x. f x = y"  | 
|
34  | 
||
35  | 
lemma myinv_f_f: "inj f ==> myinv f (f x) = x"  | 
|
36  | 
proof -  | 
|
37  | 
assume "inj f"  | 
|
38  | 
hence "(THE x'. f x' = f x) = (THE x'. x' = x)"  | 
|
39  | 
by (simp only: inj_eq)  | 
|
40  | 
also have "... = x" by (rule the_eq_trivial)  | 
|
| 11439 | 41  | 
finally show ?thesis by (unfold myinv_def)  | 
| 11436 | 42  | 
qed  | 
43  | 
||
44  | 
lemma f_myinv_f: "inj f ==> y \<in> range f ==> f (myinv f y) = y"  | 
|
45  | 
proof (unfold myinv_def)  | 
|
46  | 
assume inj: "inj f"  | 
|
47  | 
assume "y \<in> range f"  | 
|
48  | 
then obtain x where "y = f x" ..  | 
|
49  | 
hence x: "f x = y" ..  | 
|
50  | 
thus "f (THE x. f x = y) = y"  | 
|
51  | 
proof (rule theI)  | 
|
52  | 
fix x' assume "f x' = y"  | 
|
53  | 
with x have "f x' = f x" by simp  | 
|
54  | 
with inj show "x' = x" by (rule injD)  | 
|
55  | 
qed  | 
|
56  | 
qed  | 
|
57  | 
||
58  | 
hide const myinv  | 
|
59  | 
||
60  | 
||
| 11688 | 61  | 
text {* Package setup. *}
 | 
| 10402 | 62  | 
|
| 23734 | 63  | 
theorems basic_monos =  | 
| 22218 | 64  | 
subset_refl imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj  | 
| 11688 | 65  | 
Collect_mono in_mono vimage_mono  | 
66  | 
imp_conv_disj not_not de_Morgan_disj de_Morgan_conj  | 
|
67  | 
not_all not_ex  | 
|
68  | 
Ball_def Bex_def  | 
|
| 18456 | 69  | 
induct_rulify_fallback  | 
| 11688 | 70  | 
|
| 
21018
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
71  | 
use "Tools/inductive_package.ML"  | 
| 
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
72  | 
setup InductivePackage.setup  | 
| 
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
73  | 
|
| 23734 | 74  | 
theorems [mono] =  | 
| 22218 | 75  | 
imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj  | 
| 
21018
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
76  | 
imp_conv_disj not_not de_Morgan_disj de_Morgan_conj  | 
| 
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
77  | 
not_all not_ex  | 
| 
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
78  | 
Ball_def Bex_def  | 
| 
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
79  | 
induct_rulify_fallback  | 
| 
 
e6b8d6784792
Added new package for inductive definitions, moved old package
 
berghofe 
parents: 
20604 
diff
changeset
 | 
80  | 
|
| 20604 | 81  | 
lemma False_meta_all:  | 
82  | 
"Trueprop False \<equiv> (\<And>P\<Colon>bool. P)"  | 
|
83  | 
proof  | 
|
84  | 
fix P  | 
|
85  | 
assume False  | 
|
86  | 
then show P ..  | 
|
87  | 
next  | 
|
88  | 
assume "\<And>P\<Colon>bool. P"  | 
|
| 23389 | 89  | 
then show False .  | 
| 20604 | 90  | 
qed  | 
91  | 
||
92  | 
lemma not_eq_False:  | 
|
93  | 
assumes not_eq: "x \<noteq> y"  | 
|
| 22886 | 94  | 
and eq: "x \<equiv> y"  | 
| 20604 | 95  | 
shows False  | 
96  | 
using not_eq eq by auto  | 
|
97  | 
||
98  | 
lemmas not_eq_quodlibet =  | 
|
99  | 
not_eq_False [simplified False_meta_all]  | 
|
100  | 
||
| 11688 | 101  | 
|
| 12023 | 102  | 
subsection {* Inductive datatypes and primitive recursion *}
 | 
| 11688 | 103  | 
|
| 11825 | 104  | 
text {* Package setup. *}
 | 
105  | 
||
| 10402 | 106  | 
use "Tools/datatype_aux.ML"  | 
107  | 
use "Tools/datatype_prop.ML"  | 
|
108  | 
use "Tools/datatype_rep_proofs.ML"  | 
|
109  | 
use "Tools/datatype_abs_proofs.ML"  | 
|
| 22783 | 110  | 
use "Tools/datatype_case.ML"  | 
| 13469 | 111  | 
use "Tools/datatype_realizer.ML"  | 
| 19599 | 112  | 
|
113  | 
use "Tools/datatype_hooks.ML"  | 
|
114  | 
||
| 10402 | 115  | 
use "Tools/datatype_package.ML"  | 
| 7700 | 116  | 
setup DatatypePackage.setup  | 
| 10402 | 117  | 
|
| 
12437
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
118  | 
use "Tools/datatype_codegen.ML"  | 
| 
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
119  | 
setup DatatypeCodegen.setup  | 
| 
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
120  | 
|
| 13705 | 121  | 
use "Tools/inductive_realizer.ML"  | 
122  | 
setup InductiveRealizer.setup  | 
|
123  | 
||
| 
12437
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
124  | 
use "Tools/inductive_codegen.ML"  | 
| 
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
125  | 
setup InductiveCodegen.setup  | 
| 
 
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
 
berghofe 
parents: 
12023 
diff
changeset
 | 
126  | 
|
| 10402 | 127  | 
use "Tools/primrec_package.ML"  | 
| 7700 | 128  | 
|
| 23734 | 129  | 
use "Tools/inductive_set_package.ML"  | 
130  | 
setup InductiveSetPackage.setup  | 
|
131  | 
||
| 23526 | 132  | 
text{* Lambda-abstractions with pattern matching: *}
 | 
133  | 
||
134  | 
syntax  | 
|
| 23529 | 135  | 
  "_lam_pats_syntax" :: "cases_syn => 'a => 'b"               ("(%_)" 10)
 | 
| 23526 | 136  | 
syntax (xsymbols)  | 
| 23529 | 137  | 
  "_lam_pats_syntax" :: "cases_syn => 'a => 'b"               ("(\<lambda>_)" 10)
 | 
| 23526 | 138  | 
|
| 23529 | 139  | 
parse_translation (advanced) {*
 | 
140  | 
let  | 
|
141  | 
fun fun_tr ctxt [cs] =  | 
|
142  | 
let  | 
|
143  | 
val x = Free (Name.variant (add_term_free_names (cs, [])) "x", dummyT);  | 
|
| 24349 | 144  | 
val ft = DatatypeCase.case_tr true DatatypePackage.datatype_of_constr  | 
145  | 
ctxt [x, cs]  | 
|
| 23529 | 146  | 
in lambda x ft end  | 
147  | 
in [("_lam_pats_syntax", fun_tr)] end
 | 
|
| 23526 | 148  | 
*}  | 
149  | 
||
150  | 
end  |