Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in
Toggle navigation
Menu
Open sidebar
Pierre Letouzey
natded
Commits
35a678eb
Commit
35a678eb
authored
Jul 28, 2020
by
Pierre Letouzey
Browse files
raccourci : lazy_andb_iff au lieu de andb_lazy_alt + andb_true_iff
parent
4728e3c7
Changes
2
Hide whitespace changes
Inline
Side-by-side
Meta.v
View file @
35a678eb
...
@@ -1382,7 +1382,7 @@ Proof.
...
@@ -1382,7 +1382,7 @@ Proof.
-
rewrite
<-
Names
.
is_empty_spec
.
namedec
.
-
rewrite
<-
Names
.
is_empty_spec
.
namedec
.
-
rewrite
<-
Names
.
is_empty_spec
.
namedec
.
-
rewrite
<-
Names
.
is_empty_spec
.
namedec
.
-
apply
(
term_fclosed_spec
(
Fun
""
l
)).
-
apply
(
term_fclosed_spec
(
Fun
""
l
)).
-
rewrite
<-
andb_lazy_alt
,
andb_true
_iff
,
IHf1
,
IHf2
.
-
rewrite
lazy_andb
_iff
,
IHf1
,
IHf2
.
intuition
.
intuition
.
Qed
.
Qed
.
...
...
Skolem.v
View file @
35a678eb
...
@@ -208,7 +208,7 @@ Proof.
...
@@ -208,7 +208,7 @@ Proof.
{
induction
t
as
[
|
|
f
l
IH
]
using
term_ind
'
;
cbn
;
trivial
.
{
induction
t
as
[
|
|
f
l
IH
]
using
term_ind
'
;
cbn
;
trivial
.
-
unfold
BogusPoint
.
now
rewrite
<-
SO
.
-
unfold
BogusPoint
.
now
rewrite
<-
SO
.
-
destruct
(
funsymbs
sign
f
)
as
[
ar
|
]
eqn
:
E
;
try
easy
.
-
destruct
(
funsymbs
sign
f
)
as
[
ar
|
]
eqn
:
E
;
try
easy
.
rewrite
<-
andb_lazy_alt
,
andb_true
_iff
.
intros
(
_
&
F
)
genv
lenv
.
rewrite
lazy_andb
_iff
.
intros
(
_
&
F
)
genv
lenv
.
destruct
(
FU
f
)
as
[
Hf
|
Hf
].
destruct
(
FU
f
)
as
[
Hf
|
Hf
].
+
exfalso
.
now
rewrite
(
funsOk
mo
f
),
Hf
in
E
.
+
exfalso
.
now
rewrite
(
funsOk
mo
f
),
Hf
in
E
.
+
rewrite
<-
Hf
.
destruct
(
funs
mo
f
)
as
[(
p
,
q
)
|
];
auto
.
+
rewrite
<-
Hf
.
destruct
(
funs
mo
f
)
as
[(
p
,
q
)
|
];
auto
.
...
@@ -219,12 +219,12 @@ Proof.
...
@@ -219,12 +219,12 @@ Proof.
-
intuition
.
-
intuition
.
-
intuition
.
-
intuition
.
-
destruct
(
predsymbs
sign
p
)
as
[
ar
|
]
eqn
:
E
;
try
easy
.
-
destruct
(
predsymbs
sign
p
)
as
[
ar
|
]
eqn
:
E
;
try
easy
.
rewrite
<-
andb_lazy_alt
,
andb_true
_iff
.
intros
(
_
&
F
)
genv
lenv
.
rewrite
lazy_andb
_iff
.
intros
(
_
&
F
)
genv
lenv
.
rewrite
<-
PR
.
destruct
(
preds
mo
p
)
as
[(
u
,
v
)
|
];
try
easy
.
rewrite
<-
PR
.
destruct
(
preds
mo
p
)
as
[(
u
,
v
)
|
];
try
easy
.
f_equiv
.
apply
map_ext_in
.
intros
a
Ha
.
apply
Ht
.
f_equiv
.
apply
map_ext_in
.
intros
a
Ha
.
apply
Ht
.
rewrite
forallb_forall
in
F
;
auto
.
rewrite
forallb_forall
in
F
;
auto
.
-
intros
WA
genv
lenv
.
now
rewrite
IHA
.
-
intros
WA
genv
lenv
.
now
rewrite
IHA
.
-
rewrite
<-
andb_lazy_alt
,
andb_true
_iff
.
-
rewrite
lazy_andb
_iff
.
intros
(
WA1
,
WA2
)
genv
lenv
.
intros
(
WA1
,
WA2
)
genv
lenv
.
specialize
(
IHA1
WA1
genv
lenv
).
specialize
(
IHA1
WA1
genv
lenv
).
specialize
(
IHA2
WA2
genv
lenv
).
specialize
(
IHA2
WA2
genv
lenv
).
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment