| author | wenzelm | 
| Fri, 21 Mar 2014 20:33:56 +0100 | |
| changeset 56245 | 84fc7dfa3cd4 | 
| parent 53015 | a1119cf551e8 | 
| child 63167 | 0909deb8059b | 
| permissions | -rw-r--r-- | 
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 1 | theory CR | 
| 21138 | 2 | imports Lam_Funs | 
| 18106 | 3 | begin | 
| 4 | ||
| 18269 | 5 | text {* The Church-Rosser proof from Barendregt's book *}
 | 
| 6 | ||
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 7 | lemma forget: | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 8 | assumes asm: "x\<sharp>L" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 9 | shows "L[x::=P] = L" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 10 | using asm | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 11 | proof (nominal_induct L avoiding: x P rule: lam.strong_induct) | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 12 | case (Var z) | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 13 | have "x\<sharp>Var z" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 14 | thus "(Var z)[x::=P] = (Var z)" by (simp add: fresh_atm) | 
| 18106 | 15 | next | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 16 | case (App M1 M2) | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 17 | have "x\<sharp>App M1 M2" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 18 | moreover | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 19 | have ih1: "x\<sharp>M1 \<Longrightarrow> M1[x::=P] = M1" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 20 | moreover | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 21 | have ih1: "x\<sharp>M2 \<Longrightarrow> M2[x::=P] = M2" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 22 | ultimately show "(App M1 M2)[x::=P] = (App M1 M2)" by simp | 
| 18106 | 23 | next | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 24 | case (Lam z M) | 
| 23393 | 25 | have vc: "z\<sharp>x" "z\<sharp>P" by fact+ | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 26 | have ih: "x\<sharp>M \<Longrightarrow> M[x::=P] = M" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 27 | have asm: "x\<sharp>Lam [z].M" by fact | 
| 21101 | 28 | then have "x\<sharp>M" using vc by (simp add: fresh_atm abs_fresh) | 
| 29 | then have "M[x::=P] = M" using ih by simp | |
| 30 | then show "(Lam [z].M)[x::=P] = Lam [z].M" using vc by simp | |
| 18106 | 31 | qed | 
| 32 | ||
| 18378 | 33 | lemma forget_automatic: | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 34 | assumes asm: "x\<sharp>L" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 35 | shows "L[x::=P] = L" | 
| 21101 | 36 | using asm | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 37 | by (nominal_induct L avoiding: x P rule: lam.strong_induct) | 
| 21101 | 38 | (auto simp add: abs_fresh fresh_atm) | 
| 18106 | 39 | |
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 40 | lemma fresh_fact: | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 41 | fixes z::"name" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 42 | assumes asms: "z\<sharp>N" "z\<sharp>L" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 43 | shows "z\<sharp>(N[y::=L])" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 44 | using asms | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 45 | proof (nominal_induct N avoiding: z y L rule: lam.strong_induct) | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 46 | case (Var u) | 
| 23393 | 47 | have "z\<sharp>(Var u)" "z\<sharp>L" by fact+ | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 48 | thus "z\<sharp>((Var u)[y::=L])" by simp | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 49 | next | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 50 | case (App N1 N2) | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 51 | have ih1: "\<lbrakk>z\<sharp>N1; z\<sharp>L\<rbrakk> \<Longrightarrow> z\<sharp>N1[y::=L]" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 52 | moreover | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 53 | have ih2: "\<lbrakk>z\<sharp>N2; z\<sharp>L\<rbrakk> \<Longrightarrow> z\<sharp>N2[y::=L]" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 54 | moreover | 
| 23393 | 55 | have "z\<sharp>App N1 N2" "z\<sharp>L" by fact+ | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 56 | ultimately show "z\<sharp>((App N1 N2)[y::=L])" by simp | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 57 | next | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 58 | case (Lam u N1) | 
| 23393 | 59 | have vc: "u\<sharp>z" "u\<sharp>y" "u\<sharp>L" by fact+ | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 60 | have "z\<sharp>Lam [u].N1" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 61 | hence "z\<sharp>N1" using vc by (simp add: abs_fresh fresh_atm) | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 62 | moreover | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 63 | have ih: "\<lbrakk>z\<sharp>N1; z\<sharp>L\<rbrakk> \<Longrightarrow> z\<sharp>(N1[y::=L])" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 64 | moreover | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 65 | have "z\<sharp>L" by fact | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 66 | ultimately show "z\<sharp>(Lam [u].N1)[y::=L]" using vc by (simp add: abs_fresh) | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 67 | qed | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 68 | |
| 18378 | 69 | lemma fresh_fact_automatic: | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 70 | fixes z::"name" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 71 | assumes asms: "z\<sharp>N" "z\<sharp>L" | 
| 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 72 | shows "z\<sharp>(N[y::=L])" | 
| 21101 | 73 | using asms | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 74 | by (nominal_induct N avoiding: z y L rule: lam.strong_induct) | 
| 21101 | 75 | (auto simp add: abs_fresh fresh_atm) | 
| 18106 | 76 | |
| 22540 | 77 | lemma fresh_fact': | 
| 78 | fixes a::"name" | |
| 79 | assumes a: "a\<sharp>t2" | |
| 80 | shows "a\<sharp>t1[a::=t2]" | |
| 81 | using a | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 82 | by (nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) | 
| 22540 | 83 | (auto simp add: abs_fresh fresh_atm) | 
| 84 | ||
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 85 | lemma substitution_lemma: | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 86 | assumes a: "x\<noteq>y" | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 87 | and b: "x\<sharp>L" | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 88 | shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 89 | using a b | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 90 | proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct) | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 91 | case (Var z) (* case 1: Variables*) | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 92 | have "x\<noteq>y" by fact | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 93 | have "x\<sharp>L" by fact | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 94 | show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS") | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 95 | proof - | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 96 |     { (*Case 1.1*)
 | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 97 | assume "z=x" | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 98 | have "(1)": "?LHS = N[y::=L]" using `z=x` by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 99 | have "(2)": "?RHS = N[y::=L]" using `z=x` `x\<noteq>y` by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 100 | from "(1)" "(2)" have "?LHS = ?RHS" by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 101 | } | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 102 | moreover | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 103 |     { (*Case 1.2*)
 | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 104 | assume "z=y" and "z\<noteq>x" | 
| 25996 | 105 | have "(1)": "?LHS = L" using `z\<noteq>x` `z=y` by simp | 
| 106 | have "(2)": "?RHS = L[x::=N[y::=L]]" using `z=y` by simp | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 107 | have "(3)": "L[x::=N[y::=L]] = L" using `x\<sharp>L` by (simp add: forget) | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 108 | from "(1)" "(2)" "(3)" have "?LHS = ?RHS" by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 109 | } | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 110 | moreover | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 111 |     { (*Case 1.3*)
 | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 112 | assume "z\<noteq>x" and "z\<noteq>y" | 
| 25996 | 113 | have "(1)": "?LHS = Var z" using `z\<noteq>x` `z\<noteq>y` by simp | 
| 114 | have "(2)": "?RHS = Var z" using `z\<noteq>x` `z\<noteq>y` by simp | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 115 | from "(1)" "(2)" have "?LHS = ?RHS" by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 116 | } | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 117 | ultimately show "?LHS = ?RHS" by blast | 
| 18106 | 118 | qed | 
| 119 | next | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 120 | case (Lam z M1) (* case 2: lambdas *) | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 121 | have ih: "\<lbrakk>x\<noteq>y; x\<sharp>L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 122 | have "x\<noteq>y" by fact | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 123 | have "x\<sharp>L" by fact | 
| 23393 | 124 | have fs: "z\<sharp>x" "z\<sharp>y" "z\<sharp>N" "z\<sharp>L" by fact+ | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 125 | hence "z\<sharp>N[y::=L]" by (simp add: fresh_fact) | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 126 | show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") | 
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 127 | proof - | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 128 | have "?LHS = Lam [z].(M1[x::=N][y::=L])" using `z\<sharp>x` `z\<sharp>y` `z\<sharp>N` `z\<sharp>L` by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 129 | also from ih have "\<dots> = Lam [z].(M1[y::=L][x::=N[y::=L]])" using `x\<noteq>y` `x\<sharp>L` by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 130 | also have "\<dots> = (Lam [z].(M1[y::=L]))[x::=N[y::=L]]" using `z\<sharp>x` `z\<sharp>N[y::=L]` by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 131 | also have "\<dots> = ?RHS" using `z\<sharp>y` `z\<sharp>L` by simp | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 132 | finally show "?LHS = ?RHS" . | 
| 18106 | 133 | qed | 
| 134 | next | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 135 | case (App M1 M2) (* case 3: applications *) | 
| 21101 | 136 | thus "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp | 
| 18106 | 137 | qed | 
| 138 | ||
| 20955 
65a9a30b8ece
made some proof look more like the ones in Barendregt
 urbanc parents: 
20503diff
changeset | 139 | lemma substitution_lemma_automatic: | 
| 19172 
ad36a9b42cf3
made some small changes to generate nicer latex-output
 urbanc parents: 
18882diff
changeset | 140 | assumes asm: "x\<noteq>y" "x\<sharp>L" | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 141 | shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" | 
| 21101 | 142 | using asm | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 143 | by (nominal_induct M avoiding: x y N L rule: lam.strong_induct) | 
| 21101 | 144 | (auto simp add: fresh_fact forget) | 
| 18106 | 145 | |
| 146 | section {* Beta Reduction *}
 | |
| 147 | ||
| 23760 | 148 | inductive | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 149 |   "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
 | 
| 21366 | 150 | where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 151 | b1[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^sub>\<beta>(App s2 t)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 152 | | b2[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^sub>\<beta>(App t s2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 153 | | b3[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^sub>\<beta> (Lam [a].s2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 154 | | b4[intro]: "a\<sharp>s2 \<Longrightarrow> (App (Lam [a].s1) s2)\<longrightarrow>\<^sub>\<beta>(s1[a::=s2])" | 
| 22540 | 155 | |
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 156 | equivariance Beta | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 157 | |
| 22540 | 158 | nominal_inductive Beta | 
| 159 | by (simp_all add: abs_fresh fresh_fact') | |
| 18106 | 160 | |
| 23760 | 161 | inductive | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 162 |   "Beta_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80)
 | 
| 21366 | 163 | where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 164 | bs1[intro, simp]: "M \<longrightarrow>\<^sub>\<beta>\<^sup>* M" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 165 | | bs2[intro]: "\<lbrakk>M1\<longrightarrow>\<^sub>\<beta>\<^sup>* M2; M2 \<longrightarrow>\<^sub>\<beta> M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M3" | 
| 21101 | 166 | |
| 22540 | 167 | equivariance Beta_star | 
| 168 | ||
| 21101 | 169 | lemma beta_star_trans: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 170 | assumes a1: "M1\<longrightarrow>\<^sub>\<beta>\<^sup>* M2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 171 | and a2: "M2\<longrightarrow>\<^sub>\<beta>\<^sup>* M3" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 172 | shows "M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M3" | 
| 21101 | 173 | using a2 a1 | 
| 174 | by (induct) (auto) | |
| 175 | ||
| 18106 | 176 | section {* One-Reduction *}
 | 
| 177 | ||
| 23760 | 178 | inductive | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 179 |   One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
 | 
| 21366 | 180 | where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 181 | o1[intro!]: "M\<longrightarrow>\<^sub>1M" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 182 | | o2[simp,intro!]: "\<lbrakk>t1\<longrightarrow>\<^sub>1t2;s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> (App t1 s1)\<longrightarrow>\<^sub>1(App t2 s2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 183 | | o3[simp,intro!]: "s1\<longrightarrow>\<^sub>1s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^sub>1(Lam [a].s2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 184 | | o4[simp,intro!]: "\<lbrakk>a\<sharp>(s1,s2); s1\<longrightarrow>\<^sub>1s2;t1\<longrightarrow>\<^sub>1t2\<rbrakk> \<Longrightarrow> (App (Lam [a].t1) s1)\<longrightarrow>\<^sub>1(t2[a::=s2])" | 
| 22540 | 185 | |
| 22730 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 186 | equivariance One | 
| 
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
 berghofe parents: 
22542diff
changeset | 187 | |
| 22540 | 188 | nominal_inductive One | 
| 189 | by (simp_all add: abs_fresh fresh_fact') | |
| 18106 | 190 | |
| 23760 | 191 | inductive | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 192 |   "One_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80)
 | 
| 21366 | 193 | where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 194 | os1[intro, simp]: "M \<longrightarrow>\<^sub>1\<^sup>* M" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 195 | | os2[intro]: "\<lbrakk>M1\<longrightarrow>\<^sub>1\<^sup>* M2; M2 \<longrightarrow>\<^sub>1 M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^sub>1\<^sup>* M3" | 
| 21101 | 196 | |
| 22540 | 197 | equivariance One_star | 
| 18106 | 198 | |
| 21101 | 199 | lemma one_star_trans: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 200 | assumes a1: "M1\<longrightarrow>\<^sub>1\<^sup>* M2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 201 | and a2: "M2\<longrightarrow>\<^sub>1\<^sup>* M3" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 202 | shows "M1\<longrightarrow>\<^sub>1\<^sup>* M3" | 
| 21101 | 203 | using a2 a1 | 
| 204 | by (induct) (auto) | |
| 205 | ||
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 206 | lemma one_fresh_preserv: | 
| 18378 | 207 | fixes a :: "name" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 208 | assumes a: "t\<longrightarrow>\<^sub>1s" | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 209 | and b: "a\<sharp>t" | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 210 | shows "a\<sharp>s" | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 211 | using a b | 
| 18106 | 212 | proof (induct) | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 213 | case o1 thus ?case by simp | 
| 18106 | 214 | next | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 215 | case o2 thus ?case by simp | 
| 18106 | 216 | next | 
| 21101 | 217 | case (o3 s1 s2 c) | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 218 | have ih: "a\<sharp>s1 \<Longrightarrow> a\<sharp>s2" by fact | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 219 | have c: "a\<sharp>Lam [c].s1" by fact | 
| 18106 | 220 | show ?case | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 221 | proof (cases "a=c") | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 222 | assume "a=c" thus "a\<sharp>Lam [c].s2" by (simp add: abs_fresh) | 
| 18106 | 223 | next | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 224 | assume d: "a\<noteq>c" | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 225 | with c have "a\<sharp>s1" by (simp add: abs_fresh) | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 226 | hence "a\<sharp>s2" using ih by simp | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 227 | thus "a\<sharp>Lam [c].s2" using d by (simp add: abs_fresh) | 
| 18106 | 228 | qed | 
| 229 | next | |
| 22540 | 230 | case (o4 c t1 t2 s1 s2) | 
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 231 | have i1: "a\<sharp>t1 \<Longrightarrow> a\<sharp>t2" by fact | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 232 | have i2: "a\<sharp>s1 \<Longrightarrow> a\<sharp>s2" by fact | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 233 | have as: "a\<sharp>App (Lam [c].s1) t1" by fact | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 234 | hence c1: "a\<sharp>Lam [c].s1" and c2: "a\<sharp>t1" by (simp add: fresh_prod)+ | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 235 | from c2 i1 have c3: "a\<sharp>t2" by simp | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 236 | show "a\<sharp>s2[c::=t2]" | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 237 | proof (cases "a=c") | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 238 | assume "a=c" | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 239 | thus "a\<sharp>s2[c::=t2]" using c3 by (simp add: fresh_fact') | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 240 | next | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 241 | assume d1: "a\<noteq>c" | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 242 | from c1 d1 have "a\<sharp>s1" by (simp add: abs_fresh) | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 243 | hence "a\<sharp>s2" using i2 by simp | 
| 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 244 | thus "a\<sharp>s2[c::=t2]" using c3 by (simp add: fresh_fact) | 
| 18106 | 245 | qed | 
| 246 | qed | |
| 247 | ||
| 22823 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 248 | lemma one_fresh_preserv_automatic: | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 249 | fixes a :: "name" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 250 | assumes a: "t\<longrightarrow>\<^sub>1s" | 
| 22823 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 251 | and b: "a\<sharp>t" | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 252 | shows "a\<sharp>s" | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 253 | using a b | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 254 | apply(nominal_induct avoiding: a rule: One.strong_induct) | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 255 | apply(auto simp add: abs_fresh fresh_atm fresh_fact) | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 256 | done | 
| 
fa9ff469247f
tuned some proofs in CR and properly included CR_Takahashi
 urbanc parents: 
22730diff
changeset | 257 | |
| 22540 | 258 | lemma subst_rename: | 
| 259 | assumes a: "c\<sharp>t1" | |
| 260 | shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]" | |
| 261 | using a | |
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 262 | by (nominal_induct t1 avoiding: a c t2 rule: lam.strong_induct) | 
| 22540 | 263 | (auto simp add: calc_atm fresh_atm abs_fresh) | 
| 264 | ||
| 18106 | 265 | lemma one_abs: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 266 | assumes a: "Lam [a].t\<longrightarrow>\<^sub>1t'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 267 | shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^sub>1t''" | 
| 25831 | 268 | proof - | 
| 269 | have "a\<sharp>Lam [a].t" by (simp add: abs_fresh) | |
| 270 | with a have "a\<sharp>t'" by (simp add: one_fresh_preserv) | |
| 271 | with a show ?thesis | |
| 272 | by (cases rule: One.strong_cases[where a="a" and aa="a"]) | |
| 273 | (auto simp add: lam.inject abs_fresh alpha) | |
| 274 | qed | |
| 18106 | 275 | |
| 276 | lemma one_app: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 277 | assumes a: "App t1 t2 \<longrightarrow>\<^sub>1 t'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 278 | shows "(\<exists>s1 s2. t' = App s1 s2 \<and> t1 \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2) \<or> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 279 | (\<exists>a s s1 s2. t1 = Lam [a].s \<and> t' = s1[a::=s2] \<and> s \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2 \<and> a\<sharp>(t2,s2))" | 
| 25831 | 280 | using a by (erule_tac One.cases) (auto simp add: lam.inject) | 
| 18106 | 281 | |
| 282 | lemma one_red: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 283 | assumes a: "App (Lam [a].t1) t2 \<longrightarrow>\<^sub>1 M" "a\<sharp>(t2,M)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 284 | shows "(\<exists>s1 s2. M = App (Lam [a].s1) s2 \<and> t1 \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2) \<or> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 285 | (\<exists>s1 s2. M = s1[a::=s2] \<and> t1 \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2)" | 
| 25831 | 286 | using a | 
| 287 | by (cases rule: One.strong_cases [where a="a" and aa="a"]) | |
| 288 | (auto dest: one_abs simp add: lam.inject abs_fresh alpha fresh_prod) | |
| 18106 | 289 | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 290 | text {* first case in Lemma 3.2.4*}
 | 
| 18106 | 291 | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 292 | lemma one_subst_aux: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 293 | assumes a: "N\<longrightarrow>\<^sub>1N'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 294 | shows "M[x::=N] \<longrightarrow>\<^sub>1 M[x::=N']" | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 295 | using a | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 296 | proof (nominal_induct M avoiding: x N N' rule: lam.strong_induct) | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 297 | case (Var y) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 298 | thus "Var y[x::=N] \<longrightarrow>\<^sub>1 Var y[x::=N']" by (cases "x=y") auto | 
| 18106 | 299 | next | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 300 | case (App P Q) (* application case - third line *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 301 | thus "(App P Q)[x::=N] \<longrightarrow>\<^sub>1 (App P Q)[x::=N']" using o2 by simp | 
| 18106 | 302 | next | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 303 | case (Lam y P) (* abstraction case - fourth line *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 304 | thus "(Lam [y].P)[x::=N] \<longrightarrow>\<^sub>1 (Lam [y].P)[x::=N']" using o3 by simp | 
| 18106 | 305 | qed | 
| 306 | ||
| 18378 | 307 | lemma one_subst_aux_automatic: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 308 | assumes a: "N\<longrightarrow>\<^sub>1N'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 309 | shows "M[x::=N] \<longrightarrow>\<^sub>1 M[x::=N']" | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 310 | using a | 
| 26966 
071f40487734
made the naming of the induction principles consistent: weak_induct is
 urbanc parents: 
25996diff
changeset | 311 | by (nominal_induct M avoiding: x N N' rule: lam.strong_induct) | 
| 25831 | 312 | (auto simp add: fresh_prod fresh_atm) | 
| 18106 | 313 | |
| 18312 
c68296902ddb
cleaned up further the proofs (diamond still needs work);
 urbanc parents: 
18303diff
changeset | 314 | lemma one_subst: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 315 | assumes a: "M\<longrightarrow>\<^sub>1M'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 316 | and b: "N\<longrightarrow>\<^sub>1N'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 317 | shows "M[x::=N]\<longrightarrow>\<^sub>1M'[x::=N']" | 
| 18773 
0eabf66582d0
the additional freshness-condition in the one-induction
 urbanc parents: 
18659diff
changeset | 318 | using a b | 
| 22540 | 319 | proof (nominal_induct M M' avoiding: N N' x rule: One.strong_induct) | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 320 | case (o1 M) | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 321 | thus ?case by (simp add: one_subst_aux) | 
| 18106 | 322 | next | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 323 | case (o2 M1 M2 N1 N2) | 
| 18106 | 324 | thus ?case by simp | 
| 325 | next | |
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 326 | case (o3 a M1 M2) | 
| 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 327 | thus ?case by simp | 
| 18106 | 328 | next | 
| 22540 | 329 | case (o4 a N1 N2 M1 M2 N N' x) | 
| 23393 | 330 | have vc: "a\<sharp>N" "a\<sharp>N'" "a\<sharp>x" "a\<sharp>N1" "a\<sharp>N2" by fact+ | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 331 | have asm: "N\<longrightarrow>\<^sub>1N'" by fact | 
| 18106 | 332 | show ?case | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 333 | proof - | 
| 22540 | 334 | have "(App (Lam [a].M1) N1)[x::=N] = App (Lam [a].(M1[x::=N])) (N1[x::=N])" using vc by simp | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 335 | moreover have "App (Lam [a].(M1[x::=N])) (N1[x::=N]) \<longrightarrow>\<^sub>1 M2[x::=N'][a::=N2[x::=N']]" | 
| 22540 | 336 | using o4 asm by (simp add: fresh_fact) | 
| 21143 
56695d1f45cf
changed a misplaced "also" to a "moreover" (caused a loop somehow)
 urbanc parents: 
21138diff
changeset | 337 | moreover have "M2[x::=N'][a::=N2[x::=N']] = M2[a::=N2][x::=N']" | 
| 22540 | 338 | using vc by (simp add: substitution_lemma fresh_atm) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 339 | ultimately show "(App (Lam [a].M1) N1)[x::=N] \<longrightarrow>\<^sub>1 M2[a::=N2][x::=N']" by simp | 
| 18106 | 340 | qed | 
| 341 | qed | |
| 342 | ||
| 18378 | 343 | lemma one_subst_automatic: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 344 | assumes a: "M\<longrightarrow>\<^sub>1M'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 345 | and b: "N\<longrightarrow>\<^sub>1N'" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 346 | shows "M[x::=N]\<longrightarrow>\<^sub>1M'[x::=N']" | 
| 18303 
b18fabea0fd0
modified almost everything for the new nominal_induct
 urbanc parents: 
18269diff
changeset | 347 | using a b | 
| 25831 | 348 | by (nominal_induct M M' avoiding: N N' x rule: One.strong_induct) | 
| 349 | (auto simp add: one_subst_aux substitution_lemma fresh_atm fresh_fact) | |
| 18106 | 350 | |
| 351 | lemma diamond[rule_format]: | |
| 352 | fixes M :: "lam" | |
| 353 | and M1:: "lam" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 354 | assumes a: "M\<longrightarrow>\<^sub>1M1" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 355 | and b: "M\<longrightarrow>\<^sub>1M2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 356 | shows "\<exists>M3. M1\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" | 
| 18344 | 357 | using a b | 
| 22540 | 358 | proof (nominal_induct avoiding: M1 M2 rule: One.strong_induct) | 
| 18106 | 359 | case (o1 M) (* case 1 --- M1 = M *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 360 | thus "\<exists>M3. M\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" by blast | 
| 18106 | 361 | next | 
| 22540 | 362 | case (o4 x Q Q' P P') (* case 2 --- a beta-reduction occurs*) | 
| 25831 | 363 | have vc: "x\<sharp>Q" "x\<sharp>Q'" "x\<sharp>M2" by fact+ | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 364 | have i1: "\<And>M2. Q \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 365 | have i2: "\<And>M2. P \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 366 | have "App (Lam [x].P) Q \<longrightarrow>\<^sub>1 M2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 367 | hence "(\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q') \<or> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 368 | (\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q')" using vc by (simp add: one_red) | 
| 18344 | 369 | moreover (* subcase 2.1 *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 370 |   { assume "\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q'"
 | 
| 18344 | 371 | then obtain P'' and Q'' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 372 | b1: "M2=App (Lam [x].P'') Q''" and b2: "P\<longrightarrow>\<^sub>1P''" and b3: "Q\<longrightarrow>\<^sub>1Q''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 373 | from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> P''\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 374 | then obtain P''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 375 | c1: "P'\<longrightarrow>\<^sub>1P'''" and c2: "P''\<longrightarrow>\<^sub>1P'''" by force | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 376 | from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 377 | then obtain Q''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 378 | d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by force | 
| 18344 | 379 | from c1 c2 d1 d2 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 380 | have "P'[x::=Q']\<longrightarrow>\<^sub>1P'''[x::=Q'''] \<and> App (Lam [x].P'') Q'' \<longrightarrow>\<^sub>1 P'''[x::=Q''']" | 
| 22540 | 381 | using vc b3 by (auto simp add: one_subst one_fresh_preserv) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 382 | hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast | 
| 18344 | 383 | } | 
| 384 | moreover (* subcase 2.2 *) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 385 |   { assume "\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q'"
 | 
| 18344 | 386 | then obtain P'' Q'' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 387 | b1: "M2=P''[x::=Q'']" and b2: "P\<longrightarrow>\<^sub>1P''" and b3: "Q\<longrightarrow>\<^sub>1Q''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 388 | from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> P''\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 389 | then obtain P''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 390 | c1: "P'\<longrightarrow>\<^sub>1P'''" and c2: "P''\<longrightarrow>\<^sub>1P'''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 391 | from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 392 | then obtain Q''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 393 | d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by blast | 
| 18344 | 394 | from c1 c2 d1 d2 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 395 | have "P'[x::=Q']\<longrightarrow>\<^sub>1P'''[x::=Q'''] \<and> P''[x::=Q'']\<longrightarrow>\<^sub>1P'''[x::=Q''']" | 
| 18344 | 396 | by (force simp add: one_subst) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 397 | hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast | 
| 18344 | 398 | } | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 399 | ultimately show "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" by blast | 
| 18106 | 400 | next | 
| 21101 | 401 | case (o2 P P' Q Q') (* case 3 *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 402 | have i0: "P\<longrightarrow>\<^sub>1P'" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 403 | have i0': "Q\<longrightarrow>\<^sub>1Q'" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 404 | have i1: "\<And>M2. Q \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 405 | have i2: "\<And>M2. P \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 406 | assume "App P Q \<longrightarrow>\<^sub>1 M2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 407 | hence "(\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^sub>1P'' \<and> Q\<longrightarrow>\<^sub>1Q'') \<or> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 408 | (\<exists>x P' P'' Q'. P = Lam [x].P' \<and> M2 = P''[x::=Q'] \<and> P'\<longrightarrow>\<^sub>1 P'' \<and> Q\<longrightarrow>\<^sub>1Q' \<and> x\<sharp>(Q,Q'))" | 
| 18344 | 409 | by (simp add: one_app[simplified]) | 
| 410 | moreover (* subcase 3.1 *) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 411 |   { assume "\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^sub>1P'' \<and> Q\<longrightarrow>\<^sub>1Q''"
 | 
| 18344 | 412 | then obtain P'' and Q'' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 413 | b1: "M2=App P'' Q''" and b2: "P\<longrightarrow>\<^sub>1P''" and b3: "Q\<longrightarrow>\<^sub>1Q''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 414 | from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> P''\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 415 | then obtain P''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 416 | c1: "P'\<longrightarrow>\<^sub>1P'''" and c2: "P''\<longrightarrow>\<^sub>1P'''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 417 | from b3 i1 have "\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3" by simp | 
| 18344 | 418 | then obtain Q''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 419 | d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by blast | 
| 18344 | 420 | from c1 c2 d1 d2 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 421 | have "App P' Q'\<longrightarrow>\<^sub>1App P''' Q''' \<and> App P'' Q'' \<longrightarrow>\<^sub>1 App P''' Q'''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 422 | hence "\<exists>M3. App P' Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast | 
| 18344 | 423 | } | 
| 424 | moreover (* subcase 3.2 *) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 425 |   { assume "\<exists>x P1 P'' Q''. P = Lam [x].P1 \<and> M2 = P''[x::=Q''] \<and> P1\<longrightarrow>\<^sub>1 P'' \<and> Q\<longrightarrow>\<^sub>1Q'' \<and> x\<sharp>(Q,Q'')"
 | 
| 18344 | 426 | then obtain x P1 P1'' Q'' where | 
| 22540 | 427 | b0: "P = Lam [x].P1" and b1: "M2 = P1''[x::=Q'']" and | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 428 | b2: "P1\<longrightarrow>\<^sub>1P1''" and b3: "Q\<longrightarrow>\<^sub>1Q''" and vc: "x\<sharp>(Q,Q'')" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 429 | from b0 i0 have "\<exists>P1'. P'=Lam [x].P1' \<and> P1\<longrightarrow>\<^sub>1P1'" by (simp add: one_abs) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 430 | then obtain P1' where g1: "P'=Lam [x].P1'" and g2: "P1\<longrightarrow>\<^sub>1P1'" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 431 | from g1 b0 b2 i2 have "(\<exists>M3. (Lam [x].P1')\<longrightarrow>\<^sub>1M3 \<and> (Lam [x].P1'')\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 432 | then obtain P1''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 433 | c1: "(Lam [x].P1')\<longrightarrow>\<^sub>1P1'''" and c2: "(Lam [x].P1'')\<longrightarrow>\<^sub>1P1'''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 434 | from c1 have "\<exists>R1. P1'''=Lam [x].R1 \<and> P1'\<longrightarrow>\<^sub>1R1" by (simp add: one_abs) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 435 | then obtain R1 where r1: "P1'''=Lam [x].R1" and r2: "P1'\<longrightarrow>\<^sub>1R1" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 436 | from c2 have "\<exists>R2. P1'''=Lam [x].R2 \<and> P1''\<longrightarrow>\<^sub>1R2" by (simp add: one_abs) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 437 | then obtain R2 where r3: "P1'''=Lam [x].R2" and r4: "P1''\<longrightarrow>\<^sub>1R2" by blast | 
| 18344 | 438 | from r1 r3 have r5: "R1=R2" by (simp add: lam.inject alpha) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 439 | from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3)" by simp | 
| 18344 | 440 | then obtain Q''' where | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 441 | d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by blast | 
| 18344 | 442 | from g1 r2 d1 r4 r5 d2 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 443 | have "App P' Q'\<longrightarrow>\<^sub>1R1[x::=Q'''] \<and> P1''[x::=Q'']\<longrightarrow>\<^sub>1R1[x::=Q''']" | 
| 22540 | 444 | using vc i0' by (simp add: one_subst one_fresh_preserv) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 445 | hence "\<exists>M3. App P' Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast | 
| 18344 | 446 | } | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 447 | ultimately show "\<exists>M3. App P' Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" by blast | 
| 18106 | 448 | next | 
| 21101 | 449 | case (o3 P P' x) (* case 4 *) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 450 | have i1: "P\<longrightarrow>\<^sub>1P'" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 451 | have i2: "\<And>M2. P \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 452 | have "(Lam [x].P)\<longrightarrow>\<^sub>1 M2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 453 | hence "\<exists>P''. M2=Lam [x].P'' \<and> P\<longrightarrow>\<^sub>1P''" by (simp add: one_abs) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 454 | then obtain P'' where b1: "M2=Lam [x].P''" and b2: "P\<longrightarrow>\<^sub>1P''" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 455 | from i2 b1 b2 have "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^sub>1M3 \<and> (Lam [x].P'')\<longrightarrow>\<^sub>1M3" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 456 | then obtain M3 where c1: "(Lam [x].P')\<longrightarrow>\<^sub>1M3" and c2: "(Lam [x].P'')\<longrightarrow>\<^sub>1M3" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 457 | from c1 have "\<exists>R1. M3=Lam [x].R1 \<and> P'\<longrightarrow>\<^sub>1R1" by (simp add: one_abs) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 458 | then obtain R1 where r1: "M3=Lam [x].R1" and r2: "P'\<longrightarrow>\<^sub>1R1" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 459 | from c2 have "\<exists>R2. M3=Lam [x].R2 \<and> P''\<longrightarrow>\<^sub>1R2" by (simp add: one_abs) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 460 | then obtain R2 where r3: "M3=Lam [x].R2" and r4: "P''\<longrightarrow>\<^sub>1R2" by blast | 
| 18344 | 461 | from r1 r3 have r5: "R1=R2" by (simp add: lam.inject alpha) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 462 | from r2 r4 have "(Lam [x].P')\<longrightarrow>\<^sub>1(Lam [x].R1) \<and> (Lam [x].P'')\<longrightarrow>\<^sub>1(Lam [x].R2)" | 
| 18344 | 463 | by (simp add: one_subst) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 464 | thus "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 r5 by blast | 
| 18106 | 465 | qed | 
| 466 | ||
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 467 | lemma one_lam_cong: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 468 | assumes a: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 469 | shows "(Lam [a].t1)\<longrightarrow>\<^sub>\<beta>\<^sup>*(Lam [a].t2)" | 
| 18106 | 470 | using a | 
| 471 | proof induct | |
| 21101 | 472 | case bs1 thus ?case by simp | 
| 18106 | 473 | next | 
| 21101 | 474 | case (bs2 y z) | 
| 475 | thus ?case by (blast dest: b3) | |
| 18106 | 476 | qed | 
| 477 | ||
| 18378 | 478 | lemma one_app_congL: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 479 | assumes a: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 480 | shows "App t1 s\<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s" | 
| 18106 | 481 | using a | 
| 482 | proof induct | |
| 21101 | 483 | case bs1 thus ?case by simp | 
| 18106 | 484 | next | 
| 21101 | 485 | case bs2 thus ?case by (blast dest: b1) | 
| 18106 | 486 | qed | 
| 487 | ||
| 18378 | 488 | lemma one_app_congR: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 489 | assumes a: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 490 | shows "App s t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App s t2" | 
| 18106 | 491 | using a | 
| 492 | proof induct | |
| 21101 | 493 | case bs1 thus ?case by simp | 
| 18106 | 494 | next | 
| 21101 | 495 | case bs2 thus ?case by (blast dest: b2) | 
| 18106 | 496 | qed | 
| 497 | ||
| 18378 | 498 | lemma one_app_cong: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 499 | assumes a1: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 500 | and a2: "s1\<longrightarrow>\<^sub>\<beta>\<^sup>*s2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 501 | shows "App t1 s1\<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s2" | 
| 18106 | 502 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 503 | have "App t1 s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s1" using a1 by (rule one_app_congL) | 
| 18378 | 504 | moreover | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 505 | have "App t2 s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s2" using a2 by (rule one_app_congR) | 
| 21101 | 506 | ultimately show ?thesis by (rule beta_star_trans) | 
| 18106 | 507 | qed | 
| 508 | ||
| 509 | lemma one_beta_star: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 510 | assumes a: "(t1\<longrightarrow>\<^sub>1t2)" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 511 | shows "(t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2)" | 
| 18106 | 512 | using a | 
| 22540 | 513 | proof(nominal_induct rule: One.strong_induct) | 
| 18378 | 514 | case o1 thus ?case by simp | 
| 18106 | 515 | next | 
| 18378 | 516 | case o2 thus ?case by (blast intro!: one_app_cong) | 
| 18106 | 517 | next | 
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 518 | case o3 thus ?case by (blast intro!: one_lam_cong) | 
| 18106 | 519 | next | 
| 22540 | 520 | case (o4 a s1 s2 t1 t2) | 
| 23393 | 521 | have vc: "a\<sharp>s1" "a\<sharp>s2" by fact+ | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 522 | have a1: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" and a2: "s1\<longrightarrow>\<^sub>\<beta>\<^sup>*s2" by fact+ | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 523 | have c1: "(App (Lam [a].t2) s2) \<longrightarrow>\<^sub>\<beta> (t2 [a::= s2])" using vc by (simp add: b4) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 524 | from a1 a2 have c2: "App (Lam [a].t1 ) s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App (Lam [a].t2 ) s2" | 
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 525 | by (blast intro!: one_app_cong one_lam_cong) | 
| 21101 | 526 | show ?case using c2 c1 by (blast intro: beta_star_trans) | 
| 18106 | 527 | qed | 
| 528 | ||
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 529 | lemma one_star_lam_cong: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 530 | assumes a: "t1\<longrightarrow>\<^sub>1\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 531 | shows "(Lam [a].t1)\<longrightarrow>\<^sub>1\<^sup>* (Lam [a].t2)" | 
| 18106 | 532 | using a | 
| 533 | proof induct | |
| 21101 | 534 | case os1 thus ?case by simp | 
| 18106 | 535 | next | 
| 21101 | 536 | case os2 thus ?case by (blast intro: one_star_trans) | 
| 18106 | 537 | qed | 
| 538 | ||
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 539 | lemma one_star_app_congL: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 540 | assumes a: "t1\<longrightarrow>\<^sub>1\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 541 | shows "App t1 s\<longrightarrow>\<^sub>1\<^sup>* App t2 s" | 
| 18106 | 542 | using a | 
| 543 | proof induct | |
| 21101 | 544 | case os1 thus ?case by simp | 
| 18106 | 545 | next | 
| 21101 | 546 | case os2 thus ?case by (blast intro: one_star_trans) | 
| 18106 | 547 | qed | 
| 548 | ||
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 549 | lemma one_star_app_congR: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 550 | assumes a: "t1\<longrightarrow>\<^sub>1\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 551 | shows "App s t1 \<longrightarrow>\<^sub>1\<^sup>* App s t2" | 
| 18106 | 552 | using a | 
| 553 | proof induct | |
| 21101 | 554 | case os1 thus ?case by simp | 
| 18106 | 555 | next | 
| 21101 | 556 | case os2 thus ?case by (blast intro: one_star_trans) | 
| 18106 | 557 | qed | 
| 558 | ||
| 559 | lemma beta_one_star: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 560 | assumes a: "t1\<longrightarrow>\<^sub>\<beta>t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 561 | shows "t1\<longrightarrow>\<^sub>1\<^sup>*t2" | 
| 18106 | 562 | using a | 
| 22540 | 563 | proof(induct) | 
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 564 | case b1 thus ?case by (blast intro!: one_star_app_congL) | 
| 18106 | 565 | next | 
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 566 | case b2 thus ?case by (blast intro!: one_star_app_congR) | 
| 18106 | 567 | next | 
| 18882 
454d09651d1a
 - renamed some lemmas (some had names coming from ancient
 urbanc parents: 
18773diff
changeset | 568 | case b3 thus ?case by (blast intro!: one_star_lam_cong) | 
| 18106 | 569 | next | 
| 22540 | 570 | case b4 thus ?case by auto | 
| 18106 | 571 | qed | 
| 572 | ||
| 573 | lemma trans_closure: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 574 | shows "(M1\<longrightarrow>\<^sub>1\<^sup>*M2) = (M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M2)" | 
| 18106 | 575 | proof | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 576 | assume "M1 \<longrightarrow>\<^sub>1\<^sup>* M2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 577 | then show "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M2" | 
| 18106 | 578 | proof induct | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 579 | case (os1 M1) thus "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M1" by simp | 
| 18106 | 580 | next | 
| 21101 | 581 | case (os2 M1 M2 M3) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 582 | have "M2\<longrightarrow>\<^sub>1M3" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 583 | then have "M2\<longrightarrow>\<^sub>\<beta>\<^sup>*M3" by (rule one_beta_star) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 584 | moreover have "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 585 | ultimately show "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M3" by (auto intro: beta_star_trans) | 
| 18106 | 586 | qed | 
| 587 | next | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 588 | assume "M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 589 | then show "M1\<longrightarrow>\<^sub>1\<^sup>*M2" | 
| 18106 | 590 | proof induct | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 591 | case (bs1 M1) thus "M1\<longrightarrow>\<^sub>1\<^sup>*M1" by simp | 
| 18106 | 592 | next | 
| 21101 | 593 | case (bs2 M1 M2 M3) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 594 | have "M2\<longrightarrow>\<^sub>\<beta>M3" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 595 | then have "M2\<longrightarrow>\<^sub>1\<^sup>*M3" by (rule beta_one_star) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 596 | moreover have "M1\<longrightarrow>\<^sub>1\<^sup>*M2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 597 | ultimately show "M1\<longrightarrow>\<^sub>1\<^sup>*M3" by (auto intro: one_star_trans) | 
| 18106 | 598 | qed | 
| 599 | qed | |
| 600 | ||
| 601 | lemma cr_one: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 602 | assumes a: "t\<longrightarrow>\<^sub>1\<^sup>*t1" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 603 | and b: "t\<longrightarrow>\<^sub>1t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 604 | shows "\<exists>t3. t1\<longrightarrow>\<^sub>1t3 \<and> t2\<longrightarrow>\<^sub>1\<^sup>*t3" | 
| 18344 | 605 | using a b | 
| 20503 | 606 | proof (induct arbitrary: t2) | 
| 21101 | 607 | case os1 thus ?case by force | 
| 18344 | 608 | next | 
| 21101 | 609 | case (os2 t s1 s2 t2) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 610 | have b: "s1 \<longrightarrow>\<^sub>1 s2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 611 | have h: "\<And>t2. t \<longrightarrow>\<^sub>1 t2 \<Longrightarrow> (\<exists>t3. s1 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 612 | have c: "t \<longrightarrow>\<^sub>1 t2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 613 | show "\<exists>t3. s2 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3" | 
| 18344 | 614 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 615 | from c h have "\<exists>t3. s1 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 616 | then obtain t3 where c1: "s1 \<longrightarrow>\<^sub>1 t3" and c2: "t2 \<longrightarrow>\<^sub>1\<^sup>* t3" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 617 | have "\<exists>t4. s2 \<longrightarrow>\<^sub>1 t4 \<and> t3 \<longrightarrow>\<^sub>1 t4" using b c1 by (blast intro: diamond) | 
| 21101 | 618 | thus ?thesis using c2 by (blast intro: one_star_trans) | 
| 18106 | 619 | qed | 
| 620 | qed | |
| 621 | ||
| 622 | lemma cr_one_star: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 623 | assumes a: "t\<longrightarrow>\<^sub>1\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 624 | and b: "t\<longrightarrow>\<^sub>1\<^sup>*t1" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 625 | shows "\<exists>t3. t1\<longrightarrow>\<^sub>1\<^sup>*t3\<and>t2\<longrightarrow>\<^sub>1\<^sup>*t3" | 
| 21101 | 626 | using a b | 
| 627 | proof (induct arbitrary: t1) | |
| 628 | case (os1 t) then show ?case by force | |
| 18106 | 629 | next | 
| 21101 | 630 | case (os2 t s1 s2 t1) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 631 | have c: "t \<longrightarrow>\<^sub>1\<^sup>* s1" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 632 | have c': "t \<longrightarrow>\<^sub>1\<^sup>* t1" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 633 | have d: "s1 \<longrightarrow>\<^sub>1 s2" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 634 | have "t \<longrightarrow>\<^sub>1\<^sup>* t1 \<Longrightarrow> (\<exists>t3. t1 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> s1 \<longrightarrow>\<^sub>1\<^sup>* t3)" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 635 | then obtain t3 where f1: "t1 \<longrightarrow>\<^sub>1\<^sup>* t3" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 636 | and f2: "s1 \<longrightarrow>\<^sub>1\<^sup>* t3" using c' by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 637 | from cr_one d f2 have "\<exists>t4. t3\<longrightarrow>\<^sub>1t4 \<and> s2\<longrightarrow>\<^sub>1\<^sup>*t4" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 638 | then obtain t4 where g1: "t3\<longrightarrow>\<^sub>1t4" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 639 | and g2: "s2\<longrightarrow>\<^sub>1\<^sup>*t4" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 640 | have "t1\<longrightarrow>\<^sub>1\<^sup>*t4" using f1 g1 by (blast intro: one_star_trans) | 
| 18378 | 641 | thus ?case using g2 by blast | 
| 18106 | 642 | qed | 
| 643 | ||
| 644 | lemma cr_beta_star: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 645 | assumes a1: "t\<longrightarrow>\<^sub>\<beta>\<^sup>*t1" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 646 | and a2: "t\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 647 | shows "\<exists>t3. t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" | 
| 18106 | 648 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 649 | from a1 have "t\<longrightarrow>\<^sub>1\<^sup>*t1" by (simp only: trans_closure) | 
| 18378 | 650 | moreover | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 651 | from a2 have "t\<longrightarrow>\<^sub>1\<^sup>*t2" by (simp only: trans_closure) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 652 | ultimately have "\<exists>t3. t1\<longrightarrow>\<^sub>1\<^sup>*t3 \<and> t2\<longrightarrow>\<^sub>1\<^sup>*t3" by (blast intro: cr_one_star) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 653 | then obtain t3 where "t1\<longrightarrow>\<^sub>1\<^sup>*t3" and "t2\<longrightarrow>\<^sub>1\<^sup>*t3" by blast | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 654 | hence "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" and "t2\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" by (simp_all only: trans_closure) | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
41589diff
changeset | 655 | then show "\<exists>t3. t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" by blast | 
| 18106 | 656 | qed | 
| 657 | ||
| 658 | end | |
| 659 |