| 10956 |      1 | theory Tacticals = Main:
 | 
|  |      2 | 
 | 
|  |      3 | text{*REPEAT*}
 | 
| 10963 |      4 | lemma "\<lbrakk>P\<longrightarrow>Q; Q\<longrightarrow>R; R\<longrightarrow>S; P\<rbrakk> \<Longrightarrow> S"
 | 
| 10956 |      5 | apply (drule mp, assumption)
 | 
|  |      6 | apply (drule mp, assumption)
 | 
|  |      7 | apply (drule mp, assumption)
 | 
|  |      8 | apply (assumption)
 | 
|  |      9 | done
 | 
|  |     10 | 
 | 
| 10963 |     11 | lemma "\<lbrakk>P\<longrightarrow>Q; Q\<longrightarrow>R; R\<longrightarrow>S; P\<rbrakk> \<Longrightarrow> S"
 | 
| 10956 |     12 | by (drule mp, assumption)+
 | 
|  |     13 | 
 | 
|  |     14 | text{*ORELSE with REPEAT*}
 | 
| 11711 |     15 | lemma "\<lbrakk>Q\<longrightarrow>R; P\<longrightarrow>Q; x<5\<longrightarrow>P;  Suc x < 5\<rbrakk> \<Longrightarrow> R" 
 | 
| 10956 |     16 | by (drule mp, (assumption|arith))+
 | 
|  |     17 | 
 | 
|  |     18 | text{*exercise: what's going on here?*}
 | 
| 10963 |     19 | lemma "\<lbrakk>P\<and>Q\<longrightarrow>R; P\<longrightarrow>Q; P\<rbrakk> \<Longrightarrow> R"
 | 
| 12390 |     20 | by (drule mp, (intro conjI)?, assumption+)+
 | 
| 10956 |     21 | 
 | 
|  |     22 | text{*defer and prefer*}
 | 
|  |     23 | 
 | 
|  |     24 | lemma "hard \<and> (P \<or> ~P) \<and> (Q\<longrightarrow>Q)"
 | 
| 12390 |     25 | apply (intro conjI)   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
| 10956 |     26 | defer 1   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
|  |     27 | apply blast+   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
|  |     28 | oops
 | 
|  |     29 | 
 | 
|  |     30 | lemma "ok1 \<and> ok2 \<and> doubtful"
 | 
| 12390 |     31 | apply (intro conjI)   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
| 10956 |     32 | prefer 3   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
|  |     33 | oops
 | 
|  |     34 | 
 | 
| 10987 |     35 | lemma "bigsubgoal1 \<and> bigsubgoal2 \<and> bigsubgoal3 \<and> bigsubgoal4 \<and> bigsubgoal5 \<and> bigsubgoal6"
 | 
| 12390 |     36 | apply (intro conjI)   --{* @{subgoals[display,indent=0,margin=65]} *}
 | 
| 10987 |     37 | txt{* @{subgoals[display,indent=0,margin=65]} 
 | 
|  |     38 | A total of 6 subgoals...
 | 
|  |     39 | *}
 | 
| 10963 |     40 | oops
 | 
|  |     41 | 
 | 
| 10956 |     42 | 
 | 
|  |     43 | 
 | 
|  |     44 | (*needed??*)
 | 
|  |     45 | 
 | 
|  |     46 | lemma "(P\<or>Q) \<and> (R\<or>S) \<Longrightarrow> PP"
 | 
| 12390 |     47 | apply (elim conjE disjE)
 | 
| 10956 |     48 | oops
 | 
|  |     49 | 
 | 
|  |     50 | lemma "((P\<or>Q) \<and> R) \<and> (Q \<and> (P\<or>S)) \<Longrightarrow> PP"
 | 
|  |     51 | apply (elim conjE)
 | 
|  |     52 | oops
 | 
|  |     53 | 
 | 
|  |     54 | lemma "((P\<or>Q) \<and> R) \<and> (Q \<and> (P\<or>S)) \<Longrightarrow> PP"
 | 
|  |     55 | apply (erule conjE)+
 | 
|  |     56 | oops
 | 
|  |     57 | 
 | 
|  |     58 | end
 |