|
1 theory Inductive_Examples |
|
2 imports Simple_Inductive_Package |
|
3 begin |
|
4 |
|
5 section {* Transitive closure *} |
|
6 |
|
7 simple_inductive |
|
8 trcl for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
9 where |
|
10 base: "trcl r x x" |
|
11 | step: "trcl r x y \<Longrightarrow> r y z \<Longrightarrow> trcl r x z" |
|
12 |
|
13 thm trcl_def |
|
14 thm trcl.induct |
|
15 thm base |
|
16 thm step |
|
17 thm trcl.intros |
|
18 |
|
19 lemma trcl_strong_induct: |
|
20 assumes trcl: "trcl r x y" |
|
21 and I1: "\<And>x. P x x" |
|
22 and I2: "\<And>x y z. P x y \<Longrightarrow> trcl r x y \<Longrightarrow> r y z \<Longrightarrow> P x z" |
|
23 shows "P x y" |
|
24 proof - |
|
25 from trcl |
|
26 have "P x y \<and> trcl r x y" |
|
27 proof induct |
|
28 case (base x) |
|
29 from I1 and trcl.base show ?case .. |
|
30 next |
|
31 case (step x y z) |
|
32 then have "P x y" and "trcl r x y" by simp_all |
|
33 from `P x y` `trcl r x y` `r y z` have "P x z" |
|
34 by (rule I2) |
|
35 moreover from `trcl r x y` `r y z` have "trcl r x z" |
|
36 by (rule trcl.step) |
|
37 ultimately show ?case .. |
|
38 qed |
|
39 then show ?thesis .. |
|
40 qed |
|
41 |
|
42 |
|
43 section {* Even and odd numbers *} |
|
44 |
|
45 simple_inductive |
|
46 even and odd |
|
47 where |
|
48 even0: "even 0" |
|
49 | evenS: "odd n \<Longrightarrow> even (Suc n)" |
|
50 | oddS: "even n \<Longrightarrow> odd (Suc n)" |
|
51 |
|
52 thm even_def odd_def |
|
53 thm even.induct odd.induct |
|
54 thm even0 |
|
55 thm evenS |
|
56 thm oddS |
|
57 thm even_odd.intros |
|
58 |
|
59 |
|
60 section {* Accessible part *} |
|
61 |
|
62 simple_inductive |
|
63 accpart for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
64 where |
|
65 accpartI: "(\<And>y. r y x \<Longrightarrow> accpart r y) \<Longrightarrow> accpart r x" |
|
66 |
|
67 thm accpart_def |
|
68 thm accpart.induct |
|
69 thm accpartI |
|
70 |
|
71 |
|
72 section {* Accessible part in locale *} |
|
73 |
|
74 locale rel = |
|
75 fixes r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
76 |
|
77 simple_inductive (in rel) accpart' |
|
78 where |
|
79 accpartI': "\<And>x. (\<And>y. r y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x" |
|
80 |
|
81 context rel |
|
82 begin |
|
83 |
|
84 thm accpartI' |
|
85 thm accpart'.induct |
|
86 |
|
87 end |
|
88 |
|
89 thm rel.accpartI' |
|
90 thm rel.accpart'.induct |
|
91 |
|
92 end |