doc-src/TutorialI/Rules/Tacticals.thy
author nipkow
Thu, 15 Mar 2001 15:05:51 +0100
changeset 11210 33300d16a63a
parent 10987 c36733b147e8
child 11407 138919f1a135
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10956
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     1
theory Tacticals = Main:
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     2
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     3
text{*REPEAT*}
10963
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
     4
lemma "\<lbrakk>P\<longrightarrow>Q; Q\<longrightarrow>R; R\<longrightarrow>S; P\<rbrakk> \<Longrightarrow> S"
10956
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     5
apply (drule mp, assumption)
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     6
apply (drule mp, assumption)
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     7
apply (drule mp, assumption)
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     8
apply (assumption)
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
     9
done
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    10
10963
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    11
lemma "\<lbrakk>P\<longrightarrow>Q; Q\<longrightarrow>R; R\<longrightarrow>S; P\<rbrakk> \<Longrightarrow> S"
10956
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    12
by (drule mp, assumption)+
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    13
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    14
text{*ORELSE with REPEAT*}
10963
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    15
lemma "\<lbrakk>Q\<longrightarrow>R; P\<longrightarrow>Q; x<#5\<longrightarrow>P;  Suc x < #5\<rbrakk> \<Longrightarrow> R" 
10956
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    16
by (drule mp, (assumption|arith))+
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    17
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    18
text{*exercise: what's going on here?*}
10963
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    19
lemma "\<lbrakk>P\<and>Q\<longrightarrow>R; P\<longrightarrow>Q; P\<rbrakk> \<Longrightarrow> R"
10956
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    20
by (drule mp, intro?, assumption+)+
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    21
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    22
text{*defer and prefer*}
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    23
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    24
lemma "hard \<and> (P \<or> ~P) \<and> (Q\<longrightarrow>Q)"
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    25
apply intro   --{* @{subgoals[display,indent=0,margin=65]} *}
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    26
defer 1   --{* @{subgoals[display,indent=0,margin=65]} *}
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    27
apply blast+   --{* @{subgoals[display,indent=0,margin=65]} *}
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    28
oops
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    29
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    30
lemma "ok1 \<and> ok2 \<and> doubtful"
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    31
apply intro   --{* @{subgoals[display,indent=0,margin=65]} *}
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    32
prefer 3   --{* @{subgoals[display,indent=0,margin=65]} *}
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    33
oops
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    34
10987
c36733b147e8 fixed the pr example
paulson
parents: 10963
diff changeset
    35
lemma "bigsubgoal1 \<and> bigsubgoal2 \<and> bigsubgoal3 \<and> bigsubgoal4 \<and> bigsubgoal5 \<and> bigsubgoal6"
10963
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    36
apply intro   --{* @{subgoals[display,indent=0,margin=65]} *}
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    37
pr 2  
10987
c36733b147e8 fixed the pr example
paulson
parents: 10963
diff changeset
    38
txt{* @{subgoals[display,indent=0,margin=65]} 
c36733b147e8 fixed the pr example
paulson
parents: 10963
diff changeset
    39
A total of 6 subgoals...
c36733b147e8 fixed the pr example
paulson
parents: 10963
diff changeset
    40
*}
10963
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    41
oops
f2c1a280f1e3 added a "pr" example; tidied
paulson
parents: 10956
diff changeset
    42
10956
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    43
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    44
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    45
(*needed??*)
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    46
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    47
lemma "(P\<or>Q) \<and> (R\<or>S) \<Longrightarrow> PP"
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    48
apply elim
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    49
oops
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    50
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    51
lemma "((P\<or>Q) \<and> R) \<and> (Q \<and> (P\<or>S)) \<Longrightarrow> PP"
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    52
apply (elim conjE)
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    53
oops
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    54
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    55
lemma "((P\<or>Q) \<and> R) \<and> (Q \<and> (P\<or>S)) \<Longrightarrow> PP"
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    56
apply (erule conjE)+
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    57
oops
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    58
1db8b894ada0 new examples theory Rules/Tacticals.thy
paulson
parents:
diff changeset
    59
end