src/Doc/Isar_Ref/Generic.thy
changeset 59785 4e6ab5831cc0
parent 59782 5d801eff5647
child 59853 4039d8aecda4
equal deleted inserted replaced
59784:bc04a20e5a37 59785:4e6ab5831cc0
   301   @{rail \<open>
   301   @{rail \<open>
   302     (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} |
   302     (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} |
   303       @@{method frule_tac} | @@{method cut_tac}) @{syntax goal_spec}? \<newline>
   303       @@{method frule_tac} | @@{method cut_tac}) @{syntax goal_spec}? \<newline>
   304     ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} )
   304     ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} )
   305     ;
   305     ;
   306     @@{method thin_tac} @{syntax goal_spec}? @{syntax prop}
   306     dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and') @{syntax for_fixes}
   307     (@'for' (@{syntax vars} + @'and'))?
   307     ;
   308     ;
   308     @@{method thin_tac} @{syntax goal_spec}? @{syntax prop} @{syntax for_fixes}
   309     @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +)
   309     ;
   310     (@'for' (@{syntax vars} + @'and'))?
   310     @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +) @{syntax for_fixes}
   311     ;
   311     ;
   312     @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +)
   312     @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +)
   313     ;
   313     ;
   314     @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}?
   314     @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}?
   315     ;
   315     ;
   316     (@@{method tactic} | @@{method raw_tactic}) @{syntax text}
   316     (@@{method tactic} | @@{method raw_tactic}) @{syntax text}
   317     ;
       
   318 
       
   319     dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and')
       
   320     (@'for' (@{syntax vars} + @'and'))?
       
   321   \<close>}
   317   \<close>}
   322 
   318 
   323 \begin{description}
   319 \begin{description}
   324 
   320 
   325   \item @{method rule_tac} etc. do resolution of rules with explicit
   321   \item @{method rule_tac} etc. do resolution of rules with explicit