aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend/Ctyping.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2006-09-11 16:03:02 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2006-09-11 16:03:02 +0000
commitbe43363d309cea62731e04ad10dddf3ecafcacd1 (patch)
treef66c346d51df74d6b6ee34f654178a44250a33c8 /cfrontend/Ctyping.v
parent5e8237152cad0cf08d3eea0d5de8cd8bc499df69 (diff)
downloadcompcert-kvx-be43363d309cea62731e04ad10dddf3ecafcacd1.tar.gz
compcert-kvx-be43363d309cea62731e04ad10dddf3ecafcacd1.zip
Revu traitement des structures et unions recursives. Dans Cshmgen, meilleure compilation de exit_if_false.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@94 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'cfrontend/Ctyping.v')
-rw-r--r--cfrontend/Ctyping.v13
1 files changed, 9 insertions, 4 deletions
diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v
index 797033dc..0795e1b2 100644
--- a/cfrontend/Ctyping.v
+++ b/cfrontend/Ctyping.v
@@ -190,8 +190,12 @@ Fixpoint eq_type (t1 t2: type) {struct t1}: bool :=
eq_type u1 u2 && if zeq sz1 sz2 then true else false
| Tfunction args1 res1, Tfunction args2 res2 =>
eq_typelist args1 args2 && eq_type res1 res2
- | Tstruct f1, Tstruct f2 => eq_fieldlist f1 f2
- | Tunion f1, Tunion f2 => eq_fieldlist f1 f2
+ | Tstruct id1 f1, Tstruct id2 f2 =>
+ if ident_eq id1 id2 then eq_fieldlist f1 f2 else false
+ | Tunion id1 f1, Tunion id2 f2 =>
+ if ident_eq id1 id2 then eq_fieldlist f1 f2 else false
+ | Tcomp_ptr id1, Tcomp_ptr id2 =>
+ if ident_eq id1 id2 then true else false
| _, _ => false
end
@@ -239,8 +243,9 @@ Proof.
decEq; auto.
TrueInv. decEq; auto.
TrueInv. decEq; auto.
- decEq; auto.
- decEq; auto.
+ TrueInv. subst i0. decEq; auto.
+ TrueInv. subst i0. decEq; auto.
+ TrueInv. congruence.
auto.
TrueInv. decEq; auto.
auto.