diff options
author | msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2012-02-20 20:54:12 +0000 |
---|---|---|
committer | msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2012-02-20 20:54:12 +0000 |
commit | 7c4d29628a58d0192b76aebf8da29644019aee1c (patch) | |
tree | 227905d3fe5e4f75a3f144c15fb8871725c74335 | |
parent | 70621433c2c4028650e1c09e88c496f0aba67a6f (diff) |
Use a heuristic to not simplify equality hypotheses remaining after dependent induction if they
don't look really trivial. We still are making a choice though. Fixes bug #2712.
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@14988 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r-- | tactics/tactics.ml | 8 | ||||
-rw-r--r-- | test-suite/success/dependentind.v | 24 |
2 files changed, 31 insertions, 1 deletions
diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 96ddc730e..11b0b9b81 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -2503,11 +2503,17 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id gl = tclMAP (fun id -> tclTRY (generalize_dep ~with_let:true (mkVar id))) vars] gl) gl +let rec compare_upto_variables x y = + if (isVar x || isRel x) && (isVar y || isRel y) then true + else compare_constr compare_upto_variables x y + let specialize_eqs id gl = let env = pf_env gl in let ty = pf_get_hyp_typ gl id in let evars = ref (project gl) in - let unif env evars c1 c2 = Evarconv.e_conv env evars c2 c1 in + let unif env evars c1 c2 = + compare_upto_variables c1 c2 && Evarconv.e_conv env evars c1 c2 + in let rec aux in_eqs ctx acc ty = match kind_of_term ty with | Prod (na, t, b) -> diff --git a/test-suite/success/dependentind.v b/test-suite/success/dependentind.v index 70fb75888..79d12a061 100644 --- a/test-suite/success/dependentind.v +++ b/test-suite/success/dependentind.v @@ -90,6 +90,28 @@ Proof with simpl in * ; eqns ; eauto with lambda. intro. eapply app... Defined. +Lemma weakening_ctx : forall Γ Δ τ, Γ ; Δ ⊢ τ -> + forall Δ', Γ ; Δ' ; Δ ⊢ τ. +Proof with simpl in * ; eqns ; eauto with lambda. + intros Γ Δ τ H. + + dependent induction H. + + destruct Δ as [|Δ τ'']... + induction Δ'... + + destruct Δ as [|Δ τ'']... + induction Δ'... + + destruct Δ as [|Δ τ'']... + apply abs. + specialize (IHterm Γ (empty, τ))... + + apply abs. + specialize (IHterm Γ (Δ, τ'', τ))... + + intro. eapply app... +Defined. Lemma exchange : forall Γ Δ α β τ, term (Γ, α, β ; Δ) τ -> term (Γ, β, α ; Δ) τ. Proof with simpl in * ; eqns ; eauto. @@ -112,6 +134,8 @@ Proof with simpl in * ; eqns ; eauto. eapply app... Defined. + + (** Example by Andrew Kenedy, uses simplification of the first component of dependent pairs. *) Set Implicit Arguments. |