| author | wenzelm | 
| Tue, 16 May 2023 20:26:09 +0200 | |
| changeset 78069 | e5bd9b3c6f0f | 
| parent 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 67406 | 1  | 
section\<open>Examples of Intuitionistic Reasoning\<close>  | 
| 14152 | 2  | 
|
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
3  | 
theory IFOL_examples imports IFOL begin  | 
| 14152 | 4  | 
|
| 67406 | 5  | 
text\<open>Quantifier example from the book Logic and Computation\<close>  | 
| 14152 | 6  | 
lemma "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
7  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 8  | 
apply (rule impI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
9  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 10  | 
apply (rule allI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
11  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 12  | 
apply (rule exI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
13  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 14  | 
apply (erule exE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
15  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 16  | 
apply (erule allE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
17  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 69505 | 18  | 
txt\<open>Now \<open>apply assumption\<close> fails\<close>  | 
| 14152 | 19  | 
oops  | 
20  | 
||
| 67406 | 21  | 
text\<open>Trying again, with the same first two steps\<close>  | 
| 14152 | 22  | 
lemma "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
23  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 24  | 
apply (rule impI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
25  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 26  | 
apply (rule allI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
27  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 28  | 
apply (erule exE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
29  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 30  | 
apply (rule exI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
31  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 32  | 
apply (erule allE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
33  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 34  | 
apply assumption  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
35  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 36  | 
done  | 
37  | 
||
38  | 
lemma "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"  | 
|
| 69593 | 39  | 
by (tactic \<open>IntPr.fast_tac \<^context> 1\<close>)  | 
| 14152 | 40  | 
|
| 67406 | 41  | 
text\<open>Example of Dyckhoff's method\<close>  | 
| 14152 | 42  | 
lemma "~ ~ ((P-->Q) | (Q-->P))"  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
43  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 44  | 
apply (unfold not_def)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
45  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 46  | 
apply (rule impI)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
47  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 48  | 
apply (erule disj_impE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
49  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 50  | 
apply (erule imp_impE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
51  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 52  | 
apply (erule imp_impE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
67406 
diff
changeset
 | 
53  | 
  \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 14152 | 54  | 
apply assumption  | 
55  | 
apply (erule FalseE)+  | 
|
56  | 
done  | 
|
57  | 
||
58  | 
end  |