equal
deleted
inserted
replaced
1 theory Ind_General_Scheme |
1 theory Ind_General_Scheme |
2 imports Simple_Inductive_Package Ind_Prelims |
2 imports Ind_Interface |
3 begin |
3 begin |
4 |
4 |
5 (*<*) |
|
6 datatype trm = |
|
7 Var "string" |
|
8 | App "trm" "trm" |
|
9 | Lam "string" "trm" |
|
10 |
|
11 simple_inductive |
|
12 fresh :: "string \<Rightarrow> trm \<Rightarrow> bool" |
|
13 where |
|
14 fresh_var: "a\<noteq>b \<Longrightarrow> fresh a (Var b)" |
|
15 | fresh_app: "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)" |
|
16 | fresh_lam1: "fresh a (Lam a t)" |
|
17 | fresh_lam2: "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)" |
|
18 (*>*) |
|
19 |
|
20 section {* The Code in a Nutshell\label{sec:nutshell} *} |
5 section {* The Code in a Nutshell\label{sec:nutshell} *} |
|
6 |
|
7 |
21 |
8 |
22 text {* |
9 text {* |
23 The inductive package will generate the reasoning infrastructure for |
10 The inductive package will generate the reasoning infrastructure for |
24 mutually recursive predicates @{text "pred\<^isub>1\<dots>pred\<^isub>n"}. In what |
11 mutually recursive predicates @{text "pred\<^isub>1\<dots>pred\<^isub>n"}. In what |
25 follows we will have the convention that various, possibly empty collections |
12 follows we will have the convention that various, possibly empty collections |
68 @{thm [display] accpartI[no_vars]} |
55 @{thm [display] accpartI[no_vars]} |
69 |
56 |
70 has a single recursive premise that has a precondition. As usual all |
57 has a single recursive premise that has a precondition. As usual all |
71 rules are stated without the leading meta-quantification @{text "\<And>xs"}. |
58 rules are stated without the leading meta-quantification @{text "\<And>xs"}. |
72 |
59 |
73 The code of the inductive package falls roughly in tree parts: the first |
60 The output of the inductive package will be definitions for the predicates, |
74 deals with the definitions, the second with the induction principles and |
61 induction principles and introduction rules. For the definitions we need to have the |
75 the third with the introduction rules. |
62 @{text rules} in a form where the meta-quantifiers and meta-implications are |
76 |
63 replaced by their object logic equivalents. Therefore an @{text "orule"} is |
77 For the definitions we need to have the @{text rules} in a form where |
64 of the form |
78 the meta-quantifiers and meta-implications are replaced by their object |
|
79 logic equivalents. Therefore an @{text "orule"} is of the form |
|
80 |
65 |
81 @{text [display] "orule ::= \<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"} |
66 @{text [display] "orule ::= \<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"} |
82 |
67 |
83 A definition for the predicate @{text "pred"} has then the form |
68 A definition for the predicate @{text "pred"} has then the form |
84 |
69 |
109 |
94 |
110 This can be done by instantiating the @{text "\<forall>preds"}-quantification |
95 This can be done by instantiating the @{text "\<forall>preds"}-quantification |
111 with the @{text "Ps"}. Then we are done since we are left with a simple |
96 with the @{text "Ps"}. Then we are done since we are left with a simple |
112 identity. |
97 identity. |
113 |
98 |
114 Although the user declares introduction rules @{text rules}, they must |
99 Although the user declares the introduction rules @{text rules}, they must |
115 also be derived from the @{text defs}. These derivations are a bit involved. |
100 also be derived from the @{text defs}. These derivations are a bit involved. |
116 Assuming we want to prove the introduction rule |
101 Assuming we want to prove the introduction rule |
117 |
102 |
118 @{text [display] "\<And>xs. As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"} |
103 @{text [display] "\<And>xs. As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"} |
119 |
104 |
120 then we can have assumptions of the form |
105 then we have assumptions of the form |
121 |
106 |
122 \begin{isabelle} |
107 \begin{isabelle} |
123 (i)~~@{text "As"}\\ |
108 (i)~~@{text "As"}\\ |
124 (ii)~@{text "(\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>*"} |
109 (ii)~@{text "(\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>*"} |
125 \end{isabelle} |
110 \end{isabelle} |
141 @{text [display] "pred ts"} |
126 @{text [display] "pred ts"} |
142 |
127 |
143 In the last step we removed some quantifiers and moved the precondition @{text "orules"} |
128 In the last step we removed some quantifiers and moved the precondition @{text "orules"} |
144 into the assumtion. The @{text "orules"} stand for all introduction rules that are given |
129 into the assumtion. The @{text "orules"} stand for all introduction rules that are given |
145 by the user. We apply the @{text orule} that corresponds to introduction rule we are |
130 by the user. We apply the @{text orule} that corresponds to introduction rule we are |
146 proving. This introduction rule must necessarily be of the form |
131 proving. After lifting to the meta-connectives, this introduction rule must necessarily |
|
132 be of the form |
147 |
133 |
148 @{text [display] "As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"} |
134 @{text [display] "As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"} |
149 |
135 |
150 When we apply this rule we end up in the goal state where we have to prove |
136 When we apply this rule we end up in the goal state where we have to prove |
151 goals of the form |
137 goals of the form |
165 |
151 |
166 Instantiating the universal quantifiers and then resolving with the assumptions |
152 Instantiating the universal quantifiers and then resolving with the assumptions |
167 in @{text "(iii)"} gives us @{text "pred ss"}, which is the goal we are after. |
153 in @{text "(iii)"} gives us @{text "pred ss"}, which is the goal we are after. |
168 This completes the proof for introduction rules. |
154 This completes the proof for introduction rules. |
169 |
155 |
170 What remains is to implement the reasoning outlined in this section. We do this |
156 What remains is to implement in Isabelle the reasoning outlined in this section. |
171 next. For building testcases, we use the shorthands for |
157 We will describe the code in the next section. For building testcases, we use the shorthands for |
172 @{text "even/odd"}, @{term "fresh"} and @{term "accpart"} |
158 @{text "even/odd"}, @{term "fresh"} and @{term "accpart"} |
173 given in Figure~\ref{fig:shorthands}. |
159 defined in Figure~\ref{fig:shorthands}. |
174 *} |
160 *} |
175 |
161 |
176 |
162 |
177 text_raw{* |
163 text_raw{* |
178 \begin{figure}[p] |
164 \begin{figure}[p] |