src/HOL/Library/Confluent_Quotient.thy
author wenzelm
Mon, 13 Apr 2020 22:08:14 +0200
changeset 71751 abf3e80bd815
parent 71393 fce780f9c9c6
child 73398 180981b87929
permissions -rw-r--r--
tuned NEWS;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71393
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     1
theory Confluent_Quotient imports 
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     2
  Confluence
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     3
begin
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     4
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     5
section \<open>Subdistributivity for quotients via confluence\<close>
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     6
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     7
locale confluent_quotient =
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     8
  fixes R :: "'Fb \<Rightarrow> 'Fb \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
     9
    and Ea :: "'Fa \<Rightarrow> 'Fa \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    10
    and Eb :: "'Fb \<Rightarrow> 'Fb \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    11
    and Ec :: "'Fc \<Rightarrow> 'Fc \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    12
    and Eab :: "'Fab \<Rightarrow> 'Fab \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    13
    and Ebc :: "'Fbc \<Rightarrow> 'Fbc \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    14
    and \<pi>_Faba :: "'Fab \<Rightarrow> 'Fa"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    15
    and \<pi>_Fabb :: "'Fab \<Rightarrow> 'Fb"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    16
    and \<pi>_Fbcb :: "'Fbc \<Rightarrow> 'Fb"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    17
    and \<pi>_Fbcc :: "'Fbc \<Rightarrow> 'Fc"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    18
    and rel_Fab :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'Fa \<Rightarrow> 'Fb \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    19
    and rel_Fbc :: "('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> 'Fb \<Rightarrow> 'Fc \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    20
    and rel_Fac :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> 'Fa \<Rightarrow> 'Fc \<Rightarrow> bool"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    21
    and set_Fab :: "'Fab \<Rightarrow> ('a \<times> 'b) set"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    22
    and set_Fbc :: "'Fbc \<Rightarrow> ('b \<times> 'c) set"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    23
  assumes confluent: "confluentp R"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    24
    and retract1_ab: "\<And>x y. R (\<pi>_Fabb x) y \<Longrightarrow> \<exists>z. Eab x z \<and> y = \<pi>_Fabb z"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    25
    and retract1_bc: "\<And>x y. R (\<pi>_Fbcb x) y \<Longrightarrow> \<exists>z. Ebc x z \<and> y = \<pi>_Fbcb z"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    26
    and generated: "Eb \<le> equivclp R"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    27
    and set_ab: "\<And>x y. Eab x y \<Longrightarrow> set_Fab x = set_Fab y"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    28
    and set_bc: "\<And>x y. Ebc x y \<Longrightarrow> set_Fbc x = set_Fbc y"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    29
    and transp_a: "transp Ea"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    30
    and transp_c: "transp Ec"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    31
    and equivp_ab: "equivp Eab"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    32
    and equivp_bc: "equivp Ebc"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    33
    and in_rel_Fab: "\<And>A x y. rel_Fab A x y \<longleftrightarrow> (\<exists>z. z \<in> {x. set_Fab x \<subseteq> {(x, y). A x y}} \<and> \<pi>_Faba z = x \<and> \<pi>_Fabb z = y)"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    34
    and in_rel_Fbc: "\<And>B x y. rel_Fbc B x y \<longleftrightarrow> (\<exists>z. z \<in> {x. set_Fbc x \<subseteq> {(x, y). B x y}} \<and> \<pi>_Fbcb z = x \<and> \<pi>_Fbcc z = y)"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    35
    and rel_compp: "\<And>A B. rel_Fac (A OO B) = rel_Fab A OO rel_Fbc B"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    36
    and \<pi>_Faba_respect: "rel_fun Eab Ea \<pi>_Faba \<pi>_Faba"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    37
    and \<pi>_Fbcc_respect: "rel_fun Ebc Ec \<pi>_Fbcc \<pi>_Fbcc"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    38
begin
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    39
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    40
lemma retract_ab: "R\<^sup>*\<^sup>* (\<pi>_Fabb x) y \<Longrightarrow> \<exists>z. Eab x z \<and> y = \<pi>_Fabb z"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    41
  by(induction rule: rtranclp_induct)(blast dest: retract1_ab intro: equivp_transp[OF equivp_ab] equivp_reflp[OF equivp_ab])+
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    42
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    43
lemma retract_bc: "R\<^sup>*\<^sup>* (\<pi>_Fbcb x) y \<Longrightarrow> \<exists>z. Ebc x z \<and> y = \<pi>_Fbcb z"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    44
  by(induction rule: rtranclp_induct)(blast dest: retract1_bc intro: equivp_transp[OF equivp_bc] equivp_reflp[OF equivp_bc])+
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    45
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    46
lemma subdistributivity: "rel_Fab A OO Eb OO rel_Fbc B \<le> Ea OO rel_Fac (A OO B) OO Ec"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    47
proof(rule predicate2I; elim relcomppE)
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    48
  fix x y y' z
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    49
  assume "rel_Fab A x y" and "Eb y y'" and "rel_Fbc B y' z"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    50
  then obtain xy y'z
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    51
    where xy: "set_Fab xy \<subseteq> {(a, b). A a b}" "x = \<pi>_Faba xy" "y = \<pi>_Fabb xy"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    52
      and y'z: "set_Fbc y'z \<subseteq> {(a, b). B a b}" "y' = \<pi>_Fbcb y'z" "z = \<pi>_Fbcc y'z"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    53
    by(auto simp add: in_rel_Fab in_rel_Fbc)
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    54
  from \<open>Eb y y'\<close> have "equivclp R y y'" using generated by blast
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    55
  then obtain u where u: "R\<^sup>*\<^sup>* y u" "R\<^sup>*\<^sup>* y' u"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    56
    unfolding semiconfluentp_equivclp[OF confluent[THEN confluentp_imp_semiconfluentp]]
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    57
    by(auto simp add: rtranclp_conversep)
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    58
  with xy y'z obtain xy' y'z'
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    59
    where retract1: "Eab xy xy'" "\<pi>_Fabb xy' = u"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    60
      and retract2: "Ebc y'z y'z'" "\<pi>_Fbcb y'z' = u"
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    61
    by(auto dest!: retract_ab retract_bc)
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    62
  from retract1(1) xy have "Ea x (\<pi>_Faba xy')" by(auto dest: \<pi>_Faba_respect[THEN rel_funD])
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    63
  moreover have "rel_Fab A (\<pi>_Faba xy') u" using xy retract1
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    64
    by(auto simp add: in_rel_Fab dest: set_ab)
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    65
  moreover have "rel_Fbc B u (\<pi>_Fbcc y'z')" using y'z retract2
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    66
    by(auto simp add: in_rel_Fbc dest: set_bc)
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    67
  moreover have "Ec (\<pi>_Fbcc y'z') z" using retract2 y'z equivp_symp[OF equivp_bc]
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    68
    by(auto dest: \<pi>_Fbcc_respect[THEN rel_funD])
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    69
  ultimately show "(Ea OO rel_Fac (A OO B) OO Ec) x z" unfolding rel_compp by blast
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    70
qed
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    71
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    72
end
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    73
fce780f9c9c6 new examples of BNF lifting across quotients using a new theory of confluence,
traytel
parents:
diff changeset
    74
end