Skip to content

Commit

Permalink
Fixing an old pecularity of "red": head betaiota redexes are now
Browse files Browse the repository at this point in the history
reduced in order to find some head constant to reduce.

git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16337 85f007b7-540e-0410-9357-904b9bb8a0f7
  • Loading branch information
herbelin committed Mar 21, 2013
1 parent b069aa8 commit f687552
Show file tree
Hide file tree
Showing 4 changed files with 10 additions and 1 deletion.
2 changes: 2 additions & 0 deletions CHANGES
Expand Up @@ -30,6 +30,8 @@ Tactics
- Similarly, "intuition" has been made more uniform and, where it now
fails, "dintuition" can be used. (possible source of incompatibilities)
- Tactic notations can now be defined locally to a module (use "Local" prefix).
- Tactic "red" now reduces head beta-iota redexes (potential source of
rare incompatibilities).

Program

Expand Down
4 changes: 3 additions & 1 deletion doc/refman/RefMan-tac.tex
Expand Up @@ -2899,7 +2899,9 @@ \subsection{\tt red}
\tacindex{red}

This tactic applies to a goal that has the form {\tt
forall (x:T1)\dots(xk:Tk), c t1 \dots\ tn} where {\tt c} is a constant. If
forall (x:T1)\dots(xk:Tk), t} with {\tt t}
$\beta\iota\zeta$-reducing to {\tt c t1 \dots\ tn} and {\tt c} a
constant. If
{\tt c} is transparent then it replaces {\tt c} with its definition
(say {\tt t}) and then reduces {\tt (t t1 \dots\ tn)} according to
$\beta\iota\zeta$-reduction rules.
Expand Down
1 change: 1 addition & 0 deletions pretyping/tacred.ml
Expand Up @@ -783,6 +783,7 @@ and whd_construct_stack env sigma s =
let try_red_product env sigma c =
let simpfun = clos_norm_flags betaiotazeta env sigma in
let rec redrec env x =
let x = whd_betaiota sigma x in
match kind_of_term x with
| App (f,l) ->
(match kind_of_term f with
Expand Down
4 changes: 4 additions & 0 deletions tactics/tactics.ml
Expand Up @@ -438,6 +438,10 @@ let rec intro_then_gen loc name_flag move_flag force_flag dep_flag tac gl =
gl
| _ ->
if not force_flag then raise (RefinerError IntroNeedsProduct);
(* Note: red_in_concl includes betaiotazeta and this was like *)
(* this since at least V6.3 (a pity *)
(* that intro do betaiotazeta only when reduction is needed; and *)
(* probably also a pity that intro does zeta *)
try
tclTHEN try_red_in_concl
(intro_then_gen loc name_flag move_flag force_flag dep_flag tac) gl
Expand Down

0 comments on commit f687552

Please sign in to comment.