| author | wenzelm | 
| Sun, 28 Aug 2022 14:55:40 +0200 | |
| changeset 76013 | 9ac09016d036 | 
| parent 65449 | c82e63b11b8b | 
| child 76213 | e44d86131648 | 
| permissions | -rw-r--r-- | 
| 41777 | 1 | (* Title: ZF/ex/Commutation.thy | 
| 11042 | 2 | Author: Tobias Nipkow & Sidi Ould Ehmety | 
| 3 | Copyright 1995 TU Muenchen | |
| 4 | ||
| 35762 | 5 | Commutation theory for proving the Church Rosser theorem. | 
| 11042 | 6 | *) | 
| 7 | ||
| 65449 
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
 wenzelm parents: 
46822diff
changeset | 8 | theory Commutation imports ZF begin | 
| 11042 | 9 | |
| 21233 | 10 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 11 | square :: "[i, i, i, i] => o" where | 
| 21233 | 12 | "square(r,s,t,u) == | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
41777diff
changeset | 13 | (\<forall>a b. <a,b> \<in> r \<longrightarrow> (\<forall>c. <a, c> \<in> s \<longrightarrow> (\<exists>x. <b,x> \<in> t & <c,x> \<in> u)))" | 
| 11042 | 14 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 15 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 16 | commute :: "[i, i] => o" where | 
| 21233 | 17 | "commute(r,s) == square(r,s,s,r)" | 
| 11042 | 18 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 19 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 20 | diamond :: "i=>o" where | 
| 21233 | 21 | "diamond(r) == commute(r, r)" | 
| 12867 | 22 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 23 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 24 | strip :: "i=>o" where | 
| 21233 | 25 | "strip(r) == commute(r^*, r)" | 
| 12867 | 26 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 27 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 28 | Church_Rosser :: "i => o" where | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
41777diff
changeset | 29 | "Church_Rosser(r) == (\<forall>x y. <x,y> \<in> (r \<union> converse(r))^* \<longrightarrow> | 
| 21233 | 30 | (\<exists>z. <x,z> \<in> r^* & <y,z> \<in> r^*))" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 31 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 32 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 33 | confluent :: "i=>o" where | 
| 21233 | 34 | "confluent(r) == diamond(r^*)" | 
| 12867 | 35 | |
| 36 | ||
| 37 | lemma square_sym: "square(r,s,t,u) ==> square(s,r,u,t)" | |
| 21233 | 38 | unfolding square_def by blast | 
| 12867 | 39 | |
| 40 | lemma square_subset: "[| square(r,s,t,u); t \<subseteq> t' |] ==> square(r,s,t',u)" | |
| 21233 | 41 | unfolding square_def by blast | 
| 12867 | 42 | |
| 43 | ||
| 21233 | 44 | lemma square_rtrancl: | 
| 45 | "square(r,s,s,t) ==> field(s)<=field(t) ==> square(r^*,s,s,t^*)" | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 46 | apply (unfold square_def, clarify) | 
| 12867 | 47 | apply (erule rtrancl_induct) | 
| 48 | apply (blast intro: rtrancl_refl) | |
| 49 | apply (blast intro: rtrancl_into_rtrancl) | |
| 50 | done | |
| 11042 | 51 | |
| 12867 | 52 | (* A special case of square_rtrancl_on *) | 
| 21233 | 53 | lemma diamond_strip: | 
| 54 | "diamond(r) ==> strip(r)" | |
| 12867 | 55 | apply (unfold diamond_def commute_def strip_def) | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 56 | apply (rule square_rtrancl, simp_all) | 
| 12867 | 57 | done | 
| 58 | ||
| 59 | (*** commute ***) | |
| 60 | ||
| 21233 | 61 | lemma commute_sym: "commute(r,s) ==> commute(s,r)" | 
| 62 | unfolding commute_def by (blast intro: square_sym) | |
| 12867 | 63 | |
| 21233 | 64 | lemma commute_rtrancl: | 
| 65 | "commute(r,s) ==> field(r)=field(s) ==> commute(r^*,s^*)" | |
| 66 | apply (unfold commute_def) | |
| 12867 | 67 | apply (rule square_rtrancl) | 
| 21233 | 68 | apply (rule square_sym [THEN square_rtrancl, THEN square_sym]) | 
| 12867 | 69 | apply (simp_all add: rtrancl_field) | 
| 70 | done | |
| 71 | ||
| 72 | ||
| 73 | lemma confluentD: "confluent(r) ==> diamond(r^*)" | |
| 74 | by (simp add: confluent_def) | |
| 75 | ||
| 76 | lemma strip_confluent: "strip(r) ==> confluent(r)" | |
| 77 | apply (unfold strip_def confluent_def diamond_def) | |
| 78 | apply (drule commute_rtrancl) | |
| 79 | apply (simp_all add: rtrancl_field) | |
| 80 | done | |
| 81 | ||
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
41777diff
changeset | 82 | lemma commute_Un: "[| commute(r,t); commute(s,t) |] ==> commute(r \<union> s, t)" | 
| 21233 | 83 | unfolding commute_def square_def by blast | 
| 11042 | 84 | |
| 21233 | 85 | lemma diamond_Un: | 
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
41777diff
changeset | 86 | "[| diamond(r); diamond(s); commute(r, s) |] ==> diamond(r \<union> s)" | 
| 21233 | 87 | unfolding diamond_def by (blast intro: commute_Un commute_sym) | 
| 12867 | 88 | |
| 21233 | 89 | lemma diamond_confluent: | 
| 12867 | 90 | "diamond(r) ==> confluent(r)" | 
| 91 | apply (unfold diamond_def confluent_def) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 92 | apply (erule commute_rtrancl, simp) | 
| 12867 | 93 | done | 
| 94 | ||
| 21233 | 95 | lemma confluent_Un: | 
| 96 | "[| confluent(r); confluent(s); commute(r^*, s^*); | |
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
41777diff
changeset | 97 | relation(r); relation(s) |] ==> confluent(r \<union> s)" | 
| 12867 | 98 | apply (unfold confluent_def) | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 99 | apply (rule rtrancl_Un_rtrancl [THEN subst], auto) | 
| 12867 | 100 | apply (blast dest: diamond_Un intro: diamond_confluent [THEN confluentD]) | 
| 101 | done | |
| 102 | ||
| 103 | ||
| 21233 | 104 | lemma diamond_to_confluence: | 
| 12867 | 105 | "[| diamond(r); s \<subseteq> r; r<= s^* |] ==> confluent(s)" | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 106 | apply (drule rtrancl_subset [symmetric], assumption) | 
| 12867 | 107 | apply (simp_all add: confluent_def) | 
| 108 | apply (blast intro: diamond_confluent [THEN confluentD]) | |
| 109 | done | |
| 110 | ||
| 111 | ||
| 112 | (*** Church_Rosser ***) | |
| 113 | ||
| 21233 | 114 | lemma Church_Rosser1: | 
| 12867 | 115 | "Church_Rosser(r) ==> confluent(r)" | 
| 21233 | 116 | apply (unfold confluent_def Church_Rosser_def square_def | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 117 | commute_def diamond_def, auto) | 
| 12867 | 118 | apply (drule converseI) | 
| 119 | apply (simp (no_asm_use) add: rtrancl_converse [symmetric]) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 120 | apply (drule_tac x = b in spec) | 
| 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 121 | apply (drule_tac x1 = c in spec [THEN mp]) | 
| 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 122 | apply (rule_tac b = a in rtrancl_trans) | 
| 12867 | 123 | apply (blast intro: rtrancl_mono [THEN subsetD])+ | 
| 124 | done | |
| 125 | ||
| 126 | ||
| 21233 | 127 | lemma Church_Rosser2: | 
| 12867 | 128 | "confluent(r) ==> Church_Rosser(r)" | 
| 21233 | 129 | apply (unfold confluent_def Church_Rosser_def square_def | 
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 130 | commute_def diamond_def, auto) | 
| 12867 | 131 | apply (frule fieldI1) | 
| 132 | apply (simp add: rtrancl_field) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13248diff
changeset | 133 | apply (erule rtrancl_induct, auto) | 
| 12867 | 134 | apply (blast intro: rtrancl_refl) | 
| 135 | apply (blast del: rtrancl_refl intro: r_into_rtrancl rtrancl_trans)+ | |
| 136 | done | |
| 137 | ||
| 138 | ||
| 46822 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 paulson parents: 
41777diff
changeset | 139 | lemma Church_Rosser: "Church_Rosser(r) \<longleftrightarrow> confluent(r)" | 
| 21233 | 140 | by (blast intro: Church_Rosser1 Church_Rosser2) | 
| 12867 | 141 | |
| 21233 | 142 | end |