| author | paulson <lp15@cam.ac.uk> | 
| Mon, 17 Oct 2022 16:00:41 +0100 | |
| changeset 76304 | e5162a8baa24 | 
| parent 67443 | 3abf6a722518 | 
| permissions | -rw-r--r-- | 
| 16417 | 1 | theory Tacticals imports Main begin | 
| 10956 | 2 | |
| 67406 | 3 | text\<open>REPEAT\<close> | 
| 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 | ||
| 67406 | 14 | text\<open>ORELSE with REPEAT\<close> | 
| 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 | ||
| 67406 | 18 | text\<open>exercise: what's going on here?\<close> | 
| 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 | |
| 67406 | 22 | text\<open>defer and prefer\<close> | 
| 10956 | 23 | |
| 24 | lemma "hard \<and> (P \<or> ~P) \<and> (Q\<longrightarrow>Q)" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 25 | apply (intro conjI)   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 26 | defer 1   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 27 | apply blast+   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 10956 | 28 | oops | 
| 29 | ||
| 30 | lemma "ok1 \<and> ok2 \<and> doubtful" | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 31 | apply (intro conjI)   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 32 | prefer 3   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 10956 | 33 | oops | 
| 34 | ||
| 10987 | 35 | lemma "bigsubgoal1 \<and> bigsubgoal2 \<and> bigsubgoal3 \<and> bigsubgoal4 \<and> bigsubgoal5 \<and> bigsubgoal6" | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67406diff
changeset | 36 | apply (intro conjI)   \<comment> \<open>@{subgoals[display,indent=0,margin=65]}\<close>
 | 
| 67406 | 37 | txt\<open>@{subgoals[display,indent=0,margin=65]} 
 | 
| 10987 | 38 | A total of 6 subgoals... | 
| 67406 | 39 | \<close> | 
| 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 |