begin
:: deftheorem Def1 defines $^ REWRITE1:def 1 :
for p, q, b3 being FinSequence holds
( ( ( p = {} or q = {} ) implies ( b3 = p $^ q iff b3 = p ^ q ) ) & ( p = {} or q = {} or ( b3 = p $^ q iff ex i being Element of NAT ex r being FinSequence st
( len p = i + 1 & r = p | (Seg i) & b3 = r ^ q ) ) ) );
theorem
theorem Th2:
theorem
theorem
theorem
:: deftheorem Def2 defines RedSequence REWRITE1:def 2 :
for R being Relation
for b2 being FinSequence holds
( b2 is RedSequence of R iff ( len b2 > 0 & ( for i being Element of NAT st i in dom b2 & i + 1 in dom b2 holds
[(b2 . i),(b2 . (i + 1))] in R ) ) );
theorem
canceled;
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
begin
:: deftheorem Def3 defines reduces REWRITE1:def 3 :
for R being Relation
for a, b being set holds
( R reduces a,b iff ex p being RedSequence of R st
( p . 1 = a & p . (len p) = b ) );
:: deftheorem Def4 defines are_convertible_wrt REWRITE1:def 4 :
for R being Relation
for a, b being set holds
( a,b are_convertible_wrt R iff R \/ (R ~) reduces a,b );
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem
theorem Th21:
theorem Th22:
theorem Th23:
theorem
theorem Th25:
Lm5:
for R being Relation
for a, b being set st a,b are_convertible_wrt R holds
b,a are_convertible_wrt R
theorem Th26:
theorem Th27:
theorem Th28:
theorem
theorem Th30:
theorem Th31:
theorem
theorem Th33:
:: deftheorem defines is_a_normal_form_wrt REWRITE1:def 5 :
for R being Relation
for a being set holds
( a is_a_normal_form_wrt R iff for b being set holds not [a,b] in R );
theorem Th34:
theorem Th35:
definition
let R be
Relation;
let a,
b be
set ;
pred b is_a_normal_form_of a,
R means :
Def6:
(
b is_a_normal_form_wrt R &
R reduces a,
b );
pred a,
b are_convergent_wrt R means :
Def7:
ex
c being
set st
(
R reduces a,
c &
R reduces b,
c );
pred a,
b are_divergent_wrt R means :
Def8:
ex
c being
set st
(
R reduces c,
a &
R reduces c,
b );
pred a,
b are_convergent<=1_wrt R means :
Def9:
ex
c being
set st
( (
[a,c] in R or
a = c ) & (
[b,c] in R or
b = c ) );
pred a,
b are_divergent<=1_wrt R means :
Def10:
ex
c being
set st
( (
[c,a] in R or
a = c ) & (
[c,b] in R or
b = c ) );
end;
:: deftheorem Def6 defines is_a_normal_form_of REWRITE1:def 6 :
for R being Relation
for a, b being set holds
( b is_a_normal_form_of a,R iff ( b is_a_normal_form_wrt R & R reduces a,b ) );
:: deftheorem Def7 defines are_convergent_wrt REWRITE1:def 7 :
for R being Relation
for a, b being set holds
( a,b are_convergent_wrt R iff ex c being set st
( R reduces a,c & R reduces b,c ) );
:: deftheorem Def8 defines are_divergent_wrt REWRITE1:def 8 :
for R being Relation
for a, b being set holds
( a,b are_divergent_wrt R iff ex c being set st
( R reduces c,a & R reduces c,b ) );
:: deftheorem Def9 defines are_convergent<=1_wrt REWRITE1:def 9 :
for R being Relation
for a, b being set holds
( a,b are_convergent<=1_wrt R iff ex c being set st
( ( [a,c] in R or a = c ) & ( [b,c] in R or b = c ) ) );
:: deftheorem Def10 defines are_divergent<=1_wrt REWRITE1:def 10 :
for R being Relation
for a, b being set holds
( a,b are_divergent<=1_wrt R iff ex c being set st
( ( [c,a] in R or a = c ) & ( [c,b] in R or b = c ) ) );
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
theorem
theorem
theorem Th43:
theorem
theorem Th45:
theorem Th46:
:: deftheorem Def11 defines has_a_normal_form_wrt REWRITE1:def 11 :
for R being Relation
for a being set holds
( a has_a_normal_form_wrt R iff ex b being set st b is_a_normal_form_of a,R );
theorem Th47:
definition
let R be
Relation;
let a be
set ;
assume that A1:
a has_a_normal_form_wrt R
and A2:
for
b,
c being
set st
b is_a_normal_form_of a,
R &
c is_a_normal_form_of a,
R holds
b = c
;
func nf (
a,
R)
-> set means :
Def12:
it is_a_normal_form_of a,
R;
existence
ex b1 being set st b1 is_a_normal_form_of a,R
by A1, Def11;
uniqueness
for b1, b2 being set st b1 is_a_normal_form_of a,R & b2 is_a_normal_form_of a,R holds
b1 = b2
by A2;
end;
:: deftheorem Def12 defines nf REWRITE1:def 12 :
for R being Relation
for a being set st a has_a_normal_form_wrt R & ( for b, c being set st b is_a_normal_form_of a,R & c is_a_normal_form_of a,R holds
b = c ) holds
for b3 being set holds
( b3 = nf (a,R) iff b3 is_a_normal_form_of a,R );
begin
:: deftheorem Def13 defines co-well_founded REWRITE1:def 13 :
for R being Relation holds
( R is co-well_founded iff R ~ is well_founded );
:: deftheorem Def14 defines weakly-normalizing REWRITE1:def 14 :
for R being Relation holds
( R is weakly-normalizing iff for a being set st a in field R holds
a has_a_normal_form_wrt R );
:: deftheorem defines strongly-normalizing REWRITE1:def 15 :
for R being Relation holds
( R is strongly-normalizing iff for f being ManySortedSet of NAT holds
not for i being Element of NAT holds [(f . i),(f . (i + 1))] in R );
:: deftheorem Def16 defines co-well_founded REWRITE1:def 16 :
for R being Relation holds
( R is co-well_founded iff for Y being set st Y c= field R & Y <> {} holds
ex a being set st
( a in Y & ( for b being set st b in Y & a <> b holds
not [a,b] in R ) ) );
theorem
begin
definition
let R,
Q be
Relation;
pred R commutes-weakly_with Q means
for
a,
b,
c being
set st
[a,b] in R &
[a,c] in Q holds
ex
d being
set st
(
Q reduces b,
d &
R reduces c,
d );
symmetry
for R, Q being Relation st ( for a, b, c being set st [a,b] in R & [a,c] in Q holds
ex d being set st
( Q reduces b,d & R reduces c,d ) ) holds
for a, b, c being set st [a,b] in Q & [a,c] in R holds
ex d being set st
( R reduces b,d & Q reduces c,d )
pred R commutes_with Q means :
Def18:
for
a,
b,
c being
set st
R reduces a,
b &
Q reduces a,
c holds
ex
d being
set st
(
Q reduces b,
d &
R reduces c,
d );
symmetry
for R, Q being Relation st ( for a, b, c being set st R reduces a,b & Q reduces a,c holds
ex d being set st
( Q reduces b,d & R reduces c,d ) ) holds
for a, b, c being set st Q reduces a,b & R reduces a,c holds
ex d being set st
( R reduces b,d & Q reduces c,d )
end;
:: deftheorem defines commutes-weakly_with REWRITE1:def 17 :
for R, Q being Relation holds
( R commutes-weakly_with Q iff for a, b, c being set st [a,b] in R & [a,c] in Q holds
ex d being set st
( Q reduces b,d & R reduces c,d ) );
:: deftheorem Def18 defines commutes_with REWRITE1:def 18 :
for R, Q being Relation holds
( R commutes_with Q iff for a, b, c being set st R reduces a,b & Q reduces a,c holds
ex d being set st
( Q reduces b,d & R reduces c,d ) );
theorem
definition
let R be
Relation;
attr R is
with_UN_property means :
Def19:
for
a,
b being
set st
a is_a_normal_form_wrt R &
b is_a_normal_form_wrt R &
a,
b are_convertible_wrt R holds
a = b;
attr R is
with_NF_property means
for
a,
b being
set st
a is_a_normal_form_wrt R &
a,
b are_convertible_wrt R holds
R reduces b,
a;
attr R is
subcommutative means :
Def21:
for
a,
b,
c being
set st
[a,b] in R &
[a,c] in R holds
b,
c are_convergent<=1_wrt R;
attr R is
confluent means :
Def22:
for
a,
b being
set st
a,
b are_divergent_wrt R holds
a,
b are_convergent_wrt R;
attr R is
with_Church-Rosser_property means :
Def23:
for
a,
b being
set st
a,
b are_convertible_wrt R holds
a,
b are_convergent_wrt R;
attr R is
locally-confluent means :
Def24:
for
a,
b,
c being
set st
[a,b] in R &
[a,c] in R holds
b,
c are_convergent_wrt R;
end;
:: deftheorem Def19 defines with_UN_property REWRITE1:def 19 :
for R being Relation holds
( R is with_UN_property iff for a, b being set st a is_a_normal_form_wrt R & b is_a_normal_form_wrt R & a,b are_convertible_wrt R holds
a = b );
:: deftheorem defines with_NF_property REWRITE1:def 20 :
for R being Relation holds
( R is with_NF_property iff for a, b being set st a is_a_normal_form_wrt R & a,b are_convertible_wrt R holds
R reduces b,a );
:: deftheorem Def21 defines subcommutative REWRITE1:def 21 :
for R being Relation holds
( R is subcommutative iff for a, b, c being set st [a,b] in R & [a,c] in R holds
b,c are_convergent<=1_wrt R );
:: deftheorem Def22 defines confluent REWRITE1:def 22 :
for R being Relation holds
( R is confluent iff for a, b being set st a,b are_divergent_wrt R holds
a,b are_convergent_wrt R );
:: deftheorem Def23 defines with_Church-Rosser_property REWRITE1:def 23 :
for R being Relation holds
( R is with_Church-Rosser_property iff for a, b being set st a,b are_convertible_wrt R holds
a,b are_convergent_wrt R );
:: deftheorem Def24 defines locally-confluent REWRITE1:def 24 :
for R being Relation holds
( R is locally-confluent iff for a, b, c being set st [a,b] in R & [a,c] in R holds
b,c are_convergent_wrt R );
theorem Th50:
theorem
theorem Th52:
theorem
theorem Th54:
theorem Th55:
theorem Th56:
:: deftheorem DD defines complete REWRITE1:def 25 :
for R being Relation holds
( R is complete iff ( R is confluent & R is strongly-normalizing ) );
theorem
theorem
theorem
begin
definition
let R,
Q be
Relation;
pred R,
Q are_equivalent means
for
a,
b being
set holds
(
a,
b are_convertible_wrt R iff
a,
b are_convertible_wrt Q );
symmetry
for R, Q being Relation st ( for a, b being set holds
( a,b are_convertible_wrt R iff a,b are_convertible_wrt Q ) ) holds
for a, b being set holds
( a,b are_convertible_wrt Q iff a,b are_convertible_wrt R )
;
end;
:: deftheorem defines are_equivalent REWRITE1:def 26 :
for R, Q being Relation holds
( R,Q are_equivalent iff for a, b being set holds
( a,b are_convertible_wrt R iff a,b are_convertible_wrt Q ) );
:: deftheorem Def27 defines are_critical_wrt REWRITE1:def 27 :
for R being Relation
for a, b being set holds
( a,b are_critical_wrt R iff ( a,b are_divergent<=1_wrt R & not a,b are_convergent_wrt R ) );
theorem Th60:
theorem
theorem
theorem Th63:
:: deftheorem Def28 defines Completion REWRITE1:def 28 :
for R being Relation
for b2 being complete Relation holds
( b2 is Completion of R iff for a, b being set holds
( a,b are_convertible_wrt R iff a,b are_convergent_wrt b2 ) );
theorem
theorem
theorem