|
1 theory IndExamples |
|
2 imports Main |
|
3 begin |
|
4 |
|
5 section {* Transitive Closure *} |
|
6 |
|
7 text {* |
|
8 Introduction rules: |
|
9 @{term "trcl R x x"} |
|
10 @{term "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"} |
|
11 *} |
|
12 |
|
13 definition "trcl \<equiv> \<lambda>R x y. |
|
14 \<forall>P. (\<forall>x. P x x) \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y" |
|
15 |
|
16 lemma trcl_induct: |
|
17 assumes trcl: "trcl R x y" |
|
18 shows "(\<And>x. P x x) \<Longrightarrow> (\<And>x y z. R x y \<Longrightarrow> P y z \<Longrightarrow> P x z) \<Longrightarrow> P x y" |
|
19 apply (atomize (full)) |
|
20 apply (cut_tac trcl) |
|
21 apply (unfold trcl_def) |
|
22 apply (drule spec [where x=P]) |
|
23 apply assumption |
|
24 done |
|
25 |
|
26 lemma trcl_base: "trcl R x x" |
|
27 apply (unfold trcl_def) |
|
28 apply (rule allI impI)+ |
|
29 apply (drule spec) |
|
30 apply assumption |
|
31 done |
|
32 |
|
33 lemma trcl_step: "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z" |
|
34 apply (unfold trcl_def) |
|
35 apply (rule allI impI)+ |
|
36 proof - |
|
37 case goal1 |
|
38 show ?case |
|
39 apply (rule goal1(4) [rule_format]) |
|
40 apply (rule goal1(1)) |
|
41 apply (rule goal1(2) [THEN spec [where x=P], THEN mp, THEN mp, |
|
42 OF goal1(3-4)]) |
|
43 done |
|
44 qed |
|
45 |
|
46 |
|
47 |
|
48 section {* Even and Odd Numbers, Mutually Inductive *} |
|
49 |
|
50 text {* |
|
51 Introduction rules: |
|
52 @{term "even 0"} |
|
53 @{term "odd m \<Longrightarrow> even (Suc m)"} |
|
54 @{term "even m \<Longrightarrow> odd (Suc m)"} |
|
55 *} |
|
56 |
|
57 definition "even \<equiv> \<lambda>n. |
|
58 \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> P n" |
|
59 |
|
60 definition "odd \<equiv> \<lambda>n. |
|
61 \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n" |
|
62 |
|
63 lemma even_induct: |
|
64 assumes even: "even n" |
|
65 shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n" |
|
66 apply (atomize (full)) |
|
67 apply (cut_tac even) |
|
68 apply (unfold even_def) |
|
69 apply (drule spec [where x=P]) |
|
70 apply (drule spec [where x=Q]) |
|
71 apply assumption |
|
72 done |
|
73 |
|
74 lemma odd_induct: |
|
75 assumes odd: "odd n" |
|
76 shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> Q n" |
|
77 apply (atomize (full)) |
|
78 apply (cut_tac odd) |
|
79 apply (unfold odd_def) |
|
80 apply (drule spec [where x=P]) |
|
81 apply (drule spec [where x=Q]) |
|
82 apply assumption |
|
83 done |
|
84 |
|
85 lemma even0: "even 0" |
|
86 apply (unfold even_def) |
|
87 apply (rule allI impI)+ |
|
88 apply assumption |
|
89 done |
|
90 |
|
91 lemma evenS: "odd m \<Longrightarrow> even (Suc m)" |
|
92 apply (unfold odd_def even_def) |
|
93 apply (rule allI impI)+ |
|
94 proof - |
|
95 case goal1 |
|
96 show ?case |
|
97 apply (rule goal1(3) [rule_format]) |
|
98 apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q], |
|
99 THEN mp, THEN mp, THEN mp, OF goal1(2-4)]) |
|
100 done |
|
101 qed |
|
102 |
|
103 lemma oddS: "even m \<Longrightarrow> odd (Suc m)" |
|
104 apply (unfold odd_def even_def) |
|
105 apply (rule allI impI)+ |
|
106 proof - |
|
107 case goal1 |
|
108 show ?case |
|
109 apply (rule goal1(4) [rule_format]) |
|
110 apply (rule goal1(1) [THEN spec [where x=P], THEN spec [where x=Q], |
|
111 THEN mp, THEN mp, THEN mp, OF goal1(2-4)]) |
|
112 done |
|
113 qed |
|
114 |
|
115 |
|
116 |
|
117 section {* Accessible Part *} |
|
118 |
|
119 text {* |
|
120 Introduction rules: |
|
121 @{term "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"} |
|
122 *} |
|
123 |
|
124 definition "accpart \<equiv> \<lambda>R x. \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x" |
|
125 |
|
126 lemma accpart_induct: |
|
127 assumes acc: "accpart R x" |
|
128 shows "(\<And>x. (\<And>y. R y x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x" |
|
129 apply (atomize (full)) |
|
130 apply (cut_tac acc) |
|
131 apply (unfold accpart_def) |
|
132 apply (drule spec [where x=P]) |
|
133 apply assumption |
|
134 done |
|
135 |
|
136 lemma accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x" |
|
137 apply (unfold accpart_def) |
|
138 apply (rule allI impI)+ |
|
139 proof - |
|
140 case goal1 |
|
141 note goal1' = this |
|
142 show ?case |
|
143 apply (rule goal1'(2) [rule_format]) |
|
144 proof - |
|
145 case goal1 |
|
146 show ?case |
|
147 apply (rule goal1'(1) [OF goal1, THEN spec [where x=P], |
|
148 THEN mp, OF goal1'(2)]) |
|
149 done |
|
150 qed |
|
151 qed |
|
152 |
|
153 end |