| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Fri, 01 Oct 2010 15:44:50 +0900 | |
| changeset 2502 | 074304d56eb2 | 
| parent 2497 | 7f311ed9204d | 
| child 2503 | cc5d23547341 | 
| permissions | -rw-r--r-- | 
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 1 | theory Lambda | 
| 2454 
9ffee4eb1ae1
renamed NewParser to Nominal2
 Christian Urban <urbanc@in.tum.de> parents: 
2442diff
changeset | 2 | imports "../Nominal2" | 
| 1594 | 3 | begin | 
| 4 | ||
| 5 | atom_decl name | |
| 2440 
0a36825b16c1
"isabelle make test" makes all major examples....they work up to supp theorems (excluding)
 Christian Urban <urbanc@in.tum.de> parents: 
2436diff
changeset | 6 | declare [[STEPS = 100]] | 
| 1594 | 7 | |
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 8 | nominal_datatype lam = | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 9 | Var "name" | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 10 | | App "lam" "lam" | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 11 | | Lam x::"name" l::"lam" bind x in l | 
| 2431 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2425diff
changeset | 12 | |
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 13 | thm lam.distinct | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 14 | thm lam.induct | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 15 | thm lam.exhaust | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 16 | thm lam.fv_defs | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 17 | thm lam.bn_defs | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 18 | thm lam.perm_simps | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 19 | thm lam.eq_iff | 
| 2442 
1f9360daf6e1
make copies of the "old" files
 Christian Urban <urbanc@in.tum.de> parents: 
2440diff
changeset | 20 | thm lam.eq_iff[folded alphas] | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 21 | thm lam.fv_bn_eqvt | 
| 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 22 | thm lam.size_eqvt | 
| 2431 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 Christian Urban <urbanc@in.tum.de> parents: 
2425diff
changeset | 23 | |
| 2424 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 Christian Urban <urbanc@in.tum.de> parents: 
2311diff
changeset | 24 | |
| 2497 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 25 | section {* Strong Exhausts Lemma and Strong Induction Lemma via Induct_schema *}
 | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 26 | |
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 27 | lemma strong_exhaust: | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 28 | fixes c::"'a::fs" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 29 | assumes "\<And>name ca. \<lbrakk>c = ca; y = Var name\<rbrakk> \<Longrightarrow> P" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 30 | and "\<And>lam1 lam2 ca. \<lbrakk>c = ca; y = App lam1 lam2\<rbrakk> \<Longrightarrow> P" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 31 | and "\<And>name lam ca. \<lbrakk>c = ca; atom name \<sharp> ca; y = Lam name lam\<rbrakk> \<Longrightarrow> P" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 32 | shows "P" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 33 | apply(rule_tac y="y" in lam.exhaust) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 34 | apply(rule assms(1)) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 35 | apply(auto)[2] | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 36 | apply(rule assms(2)) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 37 | apply(auto)[2] | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 38 | apply(subgoal_tac "\<exists>q. (q \<bullet> (atom name)) \<sharp> c \<and> supp (Lam name lam) \<sharp>* q") | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 39 | apply(erule exE) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 40 | apply(erule conjE) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 41 | apply(rule assms(3)) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 42 | apply(rule refl) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 43 | apply(simp add: atom_eqvt) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 44 | apply(clarify) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 45 | apply(drule supp_perm_eq[symmetric]) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 46 | apply(simp) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 47 | apply(rule at_set_avoiding2_atom) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 48 | apply(simp add: finite_supp) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 49 | apply(simp add: finite_supp) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 50 | apply(simp add: lam.fresh) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 51 | done | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 52 | |
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 53 | |
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 54 | lemma strong_induct: | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 55 | fixes c::"'a::fs" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 56 | assumes a1: "\<And>name c. P c (Var name)" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 57 | and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 58 | and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 59 | shows "P c lam" | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 60 | using assms | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 61 | apply(induction_schema) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 62 | apply(rule_tac y="lam" in strong_exhaust) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 63 | apply(blast) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 64 | apply(blast) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 65 | apply(blast) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 66 | apply(relation "measure (\<lambda>(x,y). size y)") | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 67 | apply(auto)[1] | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 68 | apply(simp add: lam.size) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 69 | apply(simp add: lam.size) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 70 | apply(simp add: lam.size) | 
| 
7f311ed9204d
test with induct_schema for simpler strong_ind proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2454diff
changeset | 71 | done | 
| 2082 
0854af516f14
cleaned up a bit the examples; added equivariance to all examples
 Christian Urban <urbanc@in.tum.de> parents: 
2041diff
changeset | 72 | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 73 | section {* Strong Induction Principles*}
 | 
| 1594 | 74 | |
| 2041 
3842464ee03b
Move 2 more to NewParser
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1954diff
changeset | 75 | (* | 
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 76 | Old way of establishing strong induction | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 77 | principles by chosing a fresh name. | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 78 | *) | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 79 | (* | 
| 1594 | 80 | lemma | 
| 81 | fixes c::"'a::fs" | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 82 | assumes a1: "\<And>name c. P c (Var name)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 83 | and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 84 | and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 85 | shows "P c lam" | 
| 1594 | 86 | proof - | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 87 | have "\<And>p. P c (p \<bullet> lam)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 88 | apply(induct lam arbitrary: c rule: lam.induct) | 
| 1805 | 89 | apply(perm_simp) | 
| 1594 | 90 | apply(rule a1) | 
| 1805 | 91 | apply(perm_simp) | 
| 1594 | 92 | apply(rule a2) | 
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 93 | apply(assumption) | 
| 1594 | 94 | apply(assumption) | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 95 | apply(subgoal_tac "\<exists>new::name. (atom new) \<sharp> (c, Lam (p \<bullet> name) (p \<bullet> lam))") | 
| 1594 | 96 | defer | 
| 97 | apply(simp add: fresh_def) | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 98 | apply(rule_tac X="supp (c, Lam (p \<bullet> name) (p \<bullet> lam))" in obtain_at_base) | 
| 1594 | 99 | apply(simp add: supp_Pair finite_supp) | 
| 100 | apply(blast) | |
| 101 | apply(erule exE) | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 102 | apply(rule_tac t="p \<bullet> Lam name lam" and | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 103 | s="(((p \<bullet> name) \<leftrightarrow> new) + p) \<bullet> Lam name lam" in subst) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 104 | apply(simp del: lam.perm) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 105 | apply(subst lam.perm) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 106 | apply(subst (2) lam.perm) | 
| 1594 | 107 | apply(rule flip_fresh_fresh) | 
| 108 | apply(simp add: fresh_def) | |
| 109 | apply(simp only: supp_fn') | |
| 110 | apply(simp) | |
| 111 | apply(simp add: fresh_Pair) | |
| 112 | apply(simp) | |
| 113 | apply(rule a3) | |
| 114 | apply(simp add: fresh_Pair) | |
| 115 | apply(drule_tac x="((p \<bullet> name) \<leftrightarrow> new) + p" in meta_spec) | |
| 116 | apply(simp) | |
| 117 | done | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 118 | then have "P c (0 \<bullet> lam)" by blast | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 119 | then show "P c lam" by simp | 
| 1594 | 120 | qed | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 121 | *) | 
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 122 | (* | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 123 | New way of establishing strong induction | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 124 | principles by using a appropriate permutation. | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 125 | *) | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 126 | (* | 
| 1594 | 127 | lemma | 
| 128 | fixes c::"'a::fs" | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 129 | assumes a1: "\<And>name c. P c (Var name)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 130 | and a2: "\<And>lam1 lam2 c. \<lbrakk>\<And>d. P d lam1; \<And>d. P d lam2\<rbrakk> \<Longrightarrow> P c (App lam1 lam2)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 131 | and a3: "\<And>name lam c. \<lbrakk>atom name \<sharp> c; \<And>d. P d lam\<rbrakk> \<Longrightarrow> P c (Lam name lam)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 132 | shows "P c lam" | 
| 1594 | 133 | proof - | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 134 | have "\<And>p. P c (p \<bullet> lam)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 135 | apply(induct lam arbitrary: c rule: lam.induct) | 
| 1805 | 136 | apply(perm_simp) | 
| 1594 | 137 | apply(rule a1) | 
| 1805 | 138 | apply(perm_simp) | 
| 1594 | 139 | apply(rule a2) | 
| 140 | apply(assumption) | |
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 141 | apply(assumption) | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 142 |     apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom name}) \<sharp>* c \<and> supp (p \<bullet> Lam name lam) \<sharp>* q")
 | 
| 1594 | 143 | apply(erule exE) | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 144 | apply(rule_tac t="p \<bullet> Lam name lam" and | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 145 | s="q \<bullet> p \<bullet> Lam name lam" in subst) | 
| 1594 | 146 | defer | 
| 1805 | 147 | apply(simp) | 
| 1594 | 148 | apply(rule a3) | 
| 149 | apply(simp add: eqvts fresh_star_def) | |
| 150 | apply(drule_tac x="q + p" in meta_spec) | |
| 151 | apply(simp) | |
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 152 | apply(rule at_set_avoiding2) | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 153 | apply(simp add: finite_supp) | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 154 | apply(simp add: finite_supp) | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 155 | apply(simp add: finite_supp) | 
| 1805 | 156 | apply(perm_simp) | 
| 1797 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 157 | apply(simp add: fresh_star_def fresh_def supp_fn') | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 158 | apply(rule supp_perm_eq) | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 159 | apply(simp) | 
| 
fddb470720f1
renamed ExLam to Lambda and completed the proof of the strong ind principle; tuned paper
 Christian Urban <urbanc@in.tum.de> parents: 
1773diff
changeset | 160 | done | 
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 161 | then have "P c (0 \<bullet> lam)" by blast | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 162 | then show "P c lam" by simp | 
| 1594 | 163 | qed | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 164 | *) | 
| 1594 | 165 | |
| 1805 | 166 | section {* Typing *}
 | 
| 167 | ||
| 168 | nominal_datatype ty = | |
| 169 | TVar string | |
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 170 | | TFun ty ty | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 171 | |
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 172 | notation | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 173 |  TFun ("_ \<rightarrow> _") 
 | 
| 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 174 | |
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 175 | (* | 
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 176 | declare ty.perm[eqvt] | 
| 1805 | 177 | |
| 178 | inductive | |
| 179 | valid :: "(name \<times> ty) list \<Rightarrow> bool" | |
| 180 | where | |
| 181 | "valid []" | |
| 182 | | "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)" | |
| 183 | ||
| 1828 | 184 | inductive | 
| 185 |   typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
 | |
| 186 | where | |
| 187 | t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T" | |
| 1947 | 188 | | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2" | 
| 1828 | 189 | | t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2" | 
| 190 | ||
| 1831 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 191 | equivariance valid | 
| 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 192 | equivariance typing | 
| 1816 
56cebe7f8e24
some small tunings (incompleted work in Lambda.thy)
 Christian Urban <urbanc@in.tum.de> parents: 
1814diff
changeset | 193 | |
| 1831 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 194 | thm valid.eqvt | 
| 
16653e702d89
first working version of the automatic equivariance procedure
 Christian Urban <urbanc@in.tum.de> parents: 
1828diff
changeset | 195 | thm typing.eqvt | 
| 1811 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1810diff
changeset | 196 | thm eqvts | 
| 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1810diff
changeset | 197 | thm eqvts_raw | 
| 
ae176476b525
implemented in thmdecls the case where eqvt-lemmas are of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1810diff
changeset | 198 | |
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 199 | thm typing.induct[of "\<Gamma>" "t" "T", no_vars] | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 200 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 201 | lemma | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 202 | fixes c::"'a::fs" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 203 | assumes a: "\<Gamma> \<turnstile> t : T" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 204 | and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 205 | and a2: "\<And>\<Gamma> t1 T1 T2 t2 c. \<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<And>d. P d \<Gamma> t1 T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1; \<And>d. P d \<Gamma> t2 T1\<rbrakk> | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 206 | \<Longrightarrow> P c \<Gamma> (App t1 t2) T2" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 207 | and a3: "\<And>x \<Gamma> T1 t T2 c. \<lbrakk>atom x \<sharp> c; atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2; \<And>d. P d ((x, T1) # \<Gamma>) t T2\<rbrakk> | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 208 | \<Longrightarrow> P c \<Gamma> (Lam x t) T1 \<rightarrow> T2" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 209 | shows "P c \<Gamma> t T" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 210 | proof - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 211 | from a have "\<And>p c. P c (p \<bullet> \<Gamma>) (p \<bullet> t) (p \<bullet> T)" | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 212 | proof (induct) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 213 | case (t_Var \<Gamma> x T p c) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 214 | then show ?case | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 215 | apply - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 216 | apply(perm_strict_simp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 217 | apply(rule a1) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 218 | apply(drule_tac p="p" in permute_boolI) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 219 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 220 | apply(assumption) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 221 | apply(rotate_tac 1) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 222 | apply(drule_tac p="p" in permute_boolI) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 223 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 224 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 225 | done | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 226 | next | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 227 | case (t_App \<Gamma> t1 T1 T2 t2 p c) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 228 | then show ?case | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 229 | apply - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 230 | apply(perm_strict_simp) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 231 | apply(rule a2) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 232 | apply(drule_tac p="p" in permute_boolI) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 233 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 234 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 235 | apply(assumption) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 236 | apply(rotate_tac 2) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 237 | apply(drule_tac p="p" in permute_boolI) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 238 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 239 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 240 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 241 | done | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 242 | next | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 243 | case (t_Lam x \<Gamma> T1 t T2 p c) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 244 | then show ?case | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 245 | apply - | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 246 |       apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom x}) \<sharp>* c \<and> 
 | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 247 | supp (p \<bullet> \<Gamma>, p \<bullet> Lam x t, p \<bullet> (T1 \<rightarrow> T2)) \<sharp>* q") | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 248 | apply(erule exE) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 249 | apply(rule_tac t="p \<bullet> \<Gamma>" and s="(q + p) \<bullet> \<Gamma>" in subst) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 250 | apply(simp only: permute_plus) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 251 | apply(rule supp_perm_eq) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 252 | apply(simp add: supp_Pair fresh_star_union) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 253 | apply(rule_tac t="p \<bullet> Lam x t" and s="(q + p) \<bullet> Lam x t" in subst) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 254 | apply(simp only: permute_plus) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 255 | apply(rule supp_perm_eq) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 256 | apply(simp add: supp_Pair fresh_star_union) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 257 | apply(rule_tac t="p \<bullet> (T1 \<rightarrow> T2)" and s="(q + p) \<bullet> (T1 \<rightarrow> T2)" in subst) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 258 | apply(simp only: permute_plus) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 259 | apply(rule supp_perm_eq) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 260 | apply(simp add: supp_Pair fresh_star_union) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 261 | apply(simp (no_asm) only: eqvts) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 262 | apply(rule a3) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 263 | apply(simp only: eqvts permute_plus) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 264 | apply(simp add: fresh_star_def) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 265 | apply(drule_tac p="q + p" in permute_boolI) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 266 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 267 | apply(assumption) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 268 | apply(rotate_tac 1) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 269 | apply(drule_tac p="q + p" in permute_boolI) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 270 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 271 | apply(assumption) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 272 | apply(drule_tac x="d" in meta_spec) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 273 | apply(drule_tac x="q + p" in meta_spec) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 274 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 275 | apply(assumption) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 276 | apply(rule at_set_avoiding2) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 277 | apply(simp add: finite_supp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 278 | apply(simp add: finite_supp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 279 | apply(simp add: finite_supp) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 280 | apply(rule_tac p="-p" in permute_boolE) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 281 | apply(perm_strict_simp add: permute_minus_cancel) | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 282 | --"supplied by the user" | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 283 | apply(simp add: fresh_star_prod) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 284 | apply(simp add: fresh_star_def) | 
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 285 | sorry | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 286 | qed | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 287 | then have "P c (0 \<bullet> \<Gamma>) (0 \<bullet> t) (0 \<bullet> T)" . | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 288 | then show "P c \<Gamma> t T" by simp | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 289 | qed | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 290 | |
| 1833 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 291 | *) | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1831diff
changeset | 292 | |
| 1810 
894930834ca8
fixed bug in thmdecls with destructing Trueprop; some initial infrastructure for eqvt-theorems of the form _ ==> _
 Christian Urban <urbanc@in.tum.de> parents: 
1805diff
changeset | 293 | |
| 1800 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 294 | section {* Matching *}
 | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 295 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 296 | definition | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 297 |   MATCH :: "('c::pt \<Rightarrow> (bool * 'a::pt * 'b::pt)) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 298 | where | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 299 | "MATCH M d x \<equiv> if (\<exists>!r. \<exists>q. M q = (True, x, r)) then (THE r. \<exists>q. M q = (True, x, r)) else d" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 300 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 301 | (* | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 302 | lemma MATCH_eqvt: | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 303 | shows "p \<bullet> (MATCH M d x) = MATCH (p \<bullet> M) (p \<bullet> d) (p \<bullet> x)" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 304 | unfolding MATCH_def | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 305 | apply(perm_simp the_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 306 | apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
 | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 307 | apply(simp) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 308 | thm eqvts_raw | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 309 | apply(subst if_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 310 | apply(subst ex1_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 311 | apply(subst permute_fun_def) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 312 | apply(subst ex_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 313 | apply(subst permute_fun_def) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 314 | apply(subst eq_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 315 | apply(subst permute_fun_app_eq[where f="M"]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 316 | apply(simp only: permute_minus_cancel) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 317 | apply(subst permute_prod.simps) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 318 | apply(subst permute_prod.simps) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 319 | apply(simp only: permute_minus_cancel) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 320 | apply(simp only: permute_bool_def) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 321 | apply(simp) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 322 | apply(subst ex1_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 323 | apply(subst permute_fun_def) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 324 | apply(subst ex_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 325 | apply(subst permute_fun_def) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 326 | apply(subst eq_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 327 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 328 | apply(simp only: eqvts) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 329 | apply(simp) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 330 | apply(subgoal_tac "(p \<bullet> (\<exists>!r. \<exists>q. M q = (True, x, r))) = (\<exists>!r. \<exists>q. (p \<bullet> M) q = (True, p \<bullet> x, r))") | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 331 | apply(drule sym) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 332 | apply(simp) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 333 | apply(rule impI) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 334 | apply(simp add: perm_bool) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 335 | apply(rule trans) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 336 | apply(rule pt_the_eqvt[OF pta at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 337 | apply(assumption) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 338 | apply(simp add: pt_ex_eqvt[OF pt at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 339 | apply(simp add: pt_eq_eqvt[OF ptb at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 340 | apply(rule cheat) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 341 | apply(rule trans) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 342 | apply(rule pt_ex1_eqvt) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 343 | apply(rule pta) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 344 | apply(rule at) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 345 | apply(simp add: pt_ex_eqvt[OF pt at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 346 | apply(simp add: pt_eq_eqvt[OF ptb at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 347 | apply(subst pt_pi_rev[OF pta at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 348 | apply(subst pt_fun_app_eq[OF pt at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 349 | apply(subst pt_pi_rev[OF pt at]) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 350 | apply(simp) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 351 | done | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 352 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 353 | lemma MATCH_cng: | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 354 | assumes a: "M1 = M2" "d1 = d2" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 355 | shows "MATCH M1 d1 x = MATCH M2 d2 x" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 356 | using a by simp | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 357 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 358 | lemma MATCH_eq: | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 359 | assumes a: "t = l x" "G x" "\<And>x'. t = l x' \<Longrightarrow> G x' \<Longrightarrow> r x' = r x" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 360 | shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = r x" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 361 | using a | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 362 | unfolding MATCH_def | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 363 | apply(subst if_P) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 364 | apply(rule_tac a="r x" in ex1I) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 365 | apply(rule_tac x="x" in exI) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 366 | apply(blast) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 367 | apply(erule exE) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 368 | apply(drule_tac x="q" in meta_spec) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 369 | apply(auto)[1] | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 370 | apply(rule the_equality) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 371 | apply(blast) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 372 | apply(erule exE) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 373 | apply(drule_tac x="q" in meta_spec) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 374 | apply(auto)[1] | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 375 | done | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 376 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 377 | lemma MATCH_eq2: | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 378 | assumes a: "t = l x1 x2" "G x1 x2" "\<And>x1' x2'. t = l x1' x2' \<Longrightarrow> G x1' x2' \<Longrightarrow> r x1' x2' = r x1 x2" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 379 | shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = r x1 x2" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 380 | sorry | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 381 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 382 | lemma MATCH_neq: | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 383 | assumes a: "\<And>x. t = l x \<Longrightarrow> G x \<Longrightarrow> False" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 384 | shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = d" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 385 | using a | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 386 | unfolding MATCH_def | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 387 | apply(subst if_not_P) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 388 | apply(blast) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 389 | apply(rule refl) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 390 | done | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 391 | |
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 392 | lemma MATCH_neq2: | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 393 | assumes a: "\<And>x1 x2. t = l x1 x2 \<Longrightarrow> G x1 x2 \<Longrightarrow> False" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 394 | shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = d" | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 395 | using a | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 396 | unfolding MATCH_def | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 397 | apply(subst if_not_P) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 398 | apply(auto) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 399 | done | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 400 | *) | 
| 
78fdc6b36a1c
changed the eqvt-tac to move only outermost permutations inside; added tracing infrastructure for the eqvt-tac
 Christian Urban <urbanc@in.tum.de> parents: 
1797diff
changeset | 401 | |
| 1954 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 402 | ML {*
 | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 403 | fun mk_avoids ctxt params name set = | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 404 | let | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 405 | val (_, ctxt') = ProofContext.add_fixes | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 406 | (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt; | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 407 | fun mk s = | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 408 | let | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 409 | val t = Syntax.read_term ctxt' s; | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 410 | val t' = list_abs_free (params, t) |> | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 411 | funpow (length params) (fn Abs (_, _, t) => t) | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 412 | in (t', HOLogic.dest_setT (fastype_of t)) end | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 413 | handle TERM _ => | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 414 |         error ("Expression " ^ quote s ^ " to be avoided in case " ^
 | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 415 | quote name ^ " is not a set type"); | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 416 | fun add_set p [] = [p] | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 417 | | add_set (t, T) ((u, U) :: ps) = | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 418 | if T = U then | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 419 | let val S = HOLogic.mk_setT T | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 420 |             in (Const (@{const_name sup}, S --> S --> S) $ u $ t, T) :: ps
 | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 421 | end | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 422 | else (u, U) :: add_set (t, T) ps | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 423 | in | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 424 | (mk #> add_set) set | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 425 | end; | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 426 | *} | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 427 | |
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 428 | |
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 429 | ML {* 
 | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 430 |   writeln (commas (map (Syntax.string_of_term @{context} o fst) 
 | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 431 |     (mk_avoids @{context} [] "t_Var" "{x}" []))) 
 | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 432 | *} | 
| 
23480003f9c5
some changes to the paper
 Christian Urban <urbanc@in.tum.de> parents: 
1950diff
changeset | 433 | |
| 1947 | 434 | |
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 435 | ML {*
 | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 436 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 437 | fun prove_strong_ind (pred_name, avoids) ctxt = | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 438 | Proof.theorem NONE (K I) [] ctxt | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 439 | |
| 2169 | 440 | local structure P = Parse and K = Keyword in | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 441 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 442 | val _ = | 
| 2169 | 443 | Outer_Syntax.local_theory_to_proof "nominal_inductive" | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 444 | "proves strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 445 | (P.xname -- (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name -- | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 446 | (P.$$$ ":" |-- P.and_list1 P.term))) []) >> prove_strong_ind) | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 447 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 448 | end; | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 449 | |
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 450 | *} | 
| 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 451 | |
| 1950 
7de54c9f81ac
eliminated command so that all compiles
 Christian Urban <urbanc@in.tum.de> parents: 
1949diff
changeset | 452 | (* | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 453 | nominal_inductive typing | 
| 1950 
7de54c9f81ac
eliminated command so that all compiles
 Christian Urban <urbanc@in.tum.de> parents: 
1949diff
changeset | 454 | *) | 
| 1949 
0b692f37a771
changed theorem_i to theorem....requires new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
1947diff
changeset | 455 | |
| 2157 | 456 | (* Substitution *) | 
| 2159 
ce00205e07ab
Single variable substitution
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2158diff
changeset | 457 | |
| 2165 | 458 | primrec match_Var_raw where | 
| 459 | "match_Var_raw (Var_raw x) = Some x" | |
| 460 | | "match_Var_raw (App_raw x y) = None" | |
| 461 | | "match_Var_raw (Lam_raw n t) = None" | |
| 462 | ||
| 463 | quotient_definition | |
| 464 | "match_Var :: lam \<Rightarrow> name option" | |
| 465 | is match_Var_raw | |
| 466 | ||
| 467 | lemma [quot_respect]: "(alpha_lam_raw ===> op =) match_Var_raw match_Var_raw" | |
| 468 | apply rule | |
| 469 | apply (induct_tac a b rule: alpha_lam_raw.induct) | |
| 470 | apply simp_all | |
| 471 | done | |
| 472 | ||
| 473 | lemmas match_Var_simps = match_Var_raw.simps[quot_lifted] | |
| 474 | ||
| 475 | primrec match_App_raw where | |
| 476 | "match_App_raw (Var_raw x) = None" | |
| 477 | | "match_App_raw (App_raw x y) = Some (x, y)" | |
| 478 | | "match_App_raw (Lam_raw n t) = None" | |
| 479 | ||
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 480 | (* | 
| 2165 | 481 | quotient_definition | 
| 482 | "match_App :: lam \<Rightarrow> (lam \<times> lam) option" | |
| 483 | is match_App_raw | |
| 484 | ||
| 485 | lemma [quot_respect]: | |
| 486 | "(alpha_lam_raw ===> option_rel (prod_rel alpha_lam_raw alpha_lam_raw)) match_App_raw match_App_raw" | |
| 487 | apply (intro fun_relI) | |
| 488 | apply (induct_tac a b rule: alpha_lam_raw.induct) | |
| 489 | apply simp_all | |
| 490 | done | |
| 491 | ||
| 492 | lemmas match_App_simps = match_App_raw.simps[quot_lifted] | |
| 493 | ||
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 494 | definition new where | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 495 | "new (s :: 'a :: fs) = (THE x. \<forall>a \<in> supp s. atom x \<noteq> a)" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 496 | |
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 497 | definition | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 498 | "match_Lam (S :: 'a :: fs) t = (if (\<exists>n s. (t = Lam n s)) then | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 499 | (let z = new (S, t) in Some (z, THE s. t = Lam z s)) else None)" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 500 | |
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 501 | lemma lam_half_inj: "(Lam z s = Lam z sa) = (s = sa)" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 502 | apply auto | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 503 | apply (simp only: lam.eq_iff alphas) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 504 | apply clarify | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 505 | apply (simp add: eqvts) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 506 | sorry | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 507 | |
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 508 | lemma match_Lam_simps: | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 509 | "match_Lam S (Var n) = None" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 510 | "match_Lam S (App l r) = None" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 511 | "z = new (S, (Lam z s)) \<Longrightarrow> match_Lam S (Lam z s) = Some (z, s)" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 512 | apply (simp_all add: match_Lam_def) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 513 | apply (simp add: lam_half_inj) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 514 | apply auto | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 515 | done | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 516 | *) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 517 | (* | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 518 | lemma match_Lam_simps2: | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 519 | "atom n \<sharp> ((S :: 'a :: fs), Lam n s) \<Longrightarrow> match_Lam S (Lam n s) = Some (n, s)" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 520 | apply (rule_tac t="Lam n s" | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 521 | and s="Lam (new (S, (Lam n s))) ((n \<leftrightarrow> (new (S, (Lam n s)))) \<bullet> s)" in subst) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 522 | defer | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 523 | apply (subst match_Lam_simps(3)) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 524 | defer | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 525 | apply simp | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 526 | *) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 527 | |
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 528 | (*primrec match_Lam_raw where | 
| 2165 | 529 | "match_Lam_raw (S :: atom set) (Var_raw x) = None" | 
| 530 | | "match_Lam_raw S (App_raw x y) = None" | |
| 531 | | "match_Lam_raw S (Lam_raw n t) = (let z = new (S \<union> (fv_lam_raw t - {atom n})) in Some (z, (n \<leftrightarrow> z) \<bullet> t))"
 | |
| 532 | ||
| 533 | quotient_definition | |
| 534 | "match_Lam :: (atom set) \<Rightarrow> lam \<Rightarrow> (name \<times> lam) option" | |
| 535 | is match_Lam_raw | |
| 536 | ||
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 537 | lemma swap_fresh: | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 538 | assumes a: "fv_lam_raw t \<sharp>* p" | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 539 | shows "alpha_lam_raw (p \<bullet> t) t" | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 540 | using a apply (induct t) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 541 | apply (simp add: supp_at_base fresh_star_def) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 542 | apply (rule alpha_lam_raw.intros) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 543 | apply (metis Rep_name_inverse atom_eqvt atom_name_def fresh_perm) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 544 | apply (simp) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 545 | apply (simp only: fresh_star_union) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 546 | apply clarify | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 547 | apply (rule alpha_lam_raw.intros) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 548 | apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 549 | apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 550 | apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 551 | apply (rule alpha_lam_raw.intros) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 552 | sorry | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 553 | |
| 2165 | 554 | lemma [quot_respect]: | 
| 555 | "(op = ===> alpha_lam_raw ===> option_rel (prod_rel op = alpha_lam_raw)) match_Lam_raw match_Lam_raw" | |
| 556 | proof (intro fun_relI, clarify) | |
| 557 | fix S t s | |
| 558 | assume a: "alpha_lam_raw t s" | |
| 559 | show "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S t) (match_Lam_raw S s)" | |
| 560 | using a proof (induct t s rule: alpha_lam_raw.induct) | |
| 561 | case goal1 show ?case by simp | |
| 562 | next | |
| 563 | case goal2 show ?case by simp | |
| 564 | next | |
| 565 | case (goal3 x t y s) | |
| 566 |       then obtain p where "({atom x}, t) \<approx>gen (\<lambda>x1 x2. alpha_lam_raw x1 x2 \<and>
 | |
| 567 | option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S x1) | |
| 568 |                                (match_Lam_raw S x2)) fv_lam_raw p ({atom y}, s)" ..
 | |
| 569 | then have | |
| 570 |         c: "fv_lam_raw t - {atom x} = fv_lam_raw s - {atom y}" and
 | |
| 571 |         d: "(fv_lam_raw t - {atom x}) \<sharp>* p" and
 | |
| 572 | e: "alpha_lam_raw (p \<bullet> t) s" and | |
| 573 | f: "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S (p \<bullet> t)) (match_Lam_raw S s)" and | |
| 574 |         g: "p \<bullet> {atom x} = {atom y}" unfolding alphas(1) by - (elim conjE, assumption)+
 | |
| 575 |       let ?z = "new (S \<union> (fv_lam_raw t - {atom x}))"
 | |
| 576 |       have h: "?z = new (S \<union> (fv_lam_raw s - {atom y}))" using c by simp
 | |
| 577 | show ?case | |
| 578 | unfolding match_Lam_raw.simps Let_def option_rel.simps prod_rel.simps split_conv | |
| 579 | proof | |
| 580 |         show "?z = new (S \<union> (fv_lam_raw s - {atom y}))" by (fact h)
 | |
| 581 | next | |
| 582 | have "atom y \<sharp> p" sorry | |
| 583 | have "fv_lam_raw t \<sharp>* ((x \<leftrightarrow> y) \<bullet> p)" sorry | |
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 584 | then have "alpha_lam_raw (((x \<leftrightarrow> y) \<bullet> p) \<bullet> t) t" using swap_fresh by auto | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 585 | then have "alpha_lam_raw (p \<bullet> t) ((x \<leftrightarrow> y) \<bullet> t)" sorry | 
| 2165 | 586 | have "alpha_lam_raw t ((x \<leftrightarrow> y) \<bullet> s)" sorry | 
| 587 | then have "alpha_lam_raw ((x \<leftrightarrow> ?z) \<bullet> t) ((y \<leftrightarrow> ?z) \<bullet> s)" using eqvts(15) sorry | |
| 588 |         then show "alpha_lam_raw ((x \<leftrightarrow> new (S \<union> (fv_lam_raw t - {atom x}))) \<bullet> t)
 | |
| 589 |                   ((y \<leftrightarrow> new (S \<union> (fv_lam_raw s - {atom y}))) \<bullet> s)" unfolding h .
 | |
| 590 | qed | |
| 591 | qed | |
| 592 | qed | |
| 593 | ||
| 594 | lemmas match_Lam_simps = match_Lam_raw.simps[quot_lifted] | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 595 | *) | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 596 | (* | 
| 2165 | 597 | lemma app_some: "match_App x = Some (a, b) \<Longrightarrow> x = App a b" | 
| 598 | by (induct x rule: lam.induct) (simp_all add: match_App_simps) | |
| 599 | ||
| 600 | lemma lam_some: "match_Lam S x = Some (z, s) \<Longrightarrow> x = Lam z s \<and> atom z \<sharp> S" | |
| 601 | apply (induct x rule: lam.induct) | |
| 602 | apply (simp_all add: match_Lam_simps) | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 603 | apply (thin_tac "match_Lam S lam = Some (z, s) \<Longrightarrow> lam = Lam z s \<and> atom z \<sharp> S") | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 604 | apply (simp add: match_Lam_def) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 605 | apply (subgoal_tac "\<exists>n s. Lam name lam = Lam n s") | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 606 | prefer 2 | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 607 | apply auto[1] | 
| 2165 | 608 | apply (simp add: Let_def) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 609 | apply (thin_tac "\<exists>n s. Lam name lam = Lam n s") | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 610 | apply clarify | 
| 2165 | 611 | apply (rule conjI) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 612 | apply (rule_tac t="THE s. Lam name lam = Lam (new (S, Lam name lam)) s" and | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 613 | s="(name \<leftrightarrow> (new (S, Lam name lam))) \<bullet> lam" in subst) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 614 | defer | 
| 2165 | 615 | apply (simp add: lam.eq_iff) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 616 | apply (rule_tac x="(name \<leftrightarrow> (new (S, Lam name lam)))" in exI) | 
| 2165 | 617 | apply (simp add: alphas) | 
| 618 | apply (simp add: eqvts) | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 619 | apply (rule conjI) | 
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 620 | sorry | 
| 2165 | 621 | |
| 622 | function subst where | |
| 623 | "subst v s t = ( | |
| 624 | case match_Var t of Some n \<Rightarrow> if n = v then s else Var n | None \<Rightarrow> | |
| 625 | case match_App t of Some (l, r) \<Rightarrow> App (subst v s l) (subst v s r) | None \<Rightarrow> | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 626 | case match_Lam (v,s) t of Some (n, t) \<Rightarrow> Lam n (subst v s t) | None \<Rightarrow> undefined)" | 
| 2165 | 627 | by pat_completeness auto | 
| 628 | ||
| 629 | termination apply (relation "measure (\<lambda>(_, _, t). size t)") | |
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 630 | apply auto[1] | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 631 | apply (case_tac a) apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 632 | apply (frule lam_some) apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 633 | apply (case_tac a) apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 634 | apply (frule app_some) apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 635 | apply (case_tac a) apply simp | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 636 | apply (frule app_some) apply simp | 
| 2165 | 637 | done | 
| 638 | ||
| 639 | lemmas lam_exhaust = lam_raw.exhaust[quot_lifted] | |
| 640 | ||
| 641 | lemma subst_eqvt: | |
| 642 | "p \<bullet> (subst v s t) = subst (p \<bullet> v) (p \<bullet> s) (p \<bullet> t)" | |
| 643 | proof (induct v s t rule: subst.induct) | |
| 644 | case (1 v s t) | |
| 645 | show ?case proof (cases t rule: lam_exhaust) | |
| 646 | fix n | |
| 647 | assume "t = Var n" | |
| 648 | then show ?thesis by (simp add: match_Var_simps) | |
| 649 | next | |
| 650 | fix l r | |
| 651 | assume "t = App l r" | |
| 652 | then show ?thesis | |
| 653 | apply (simp only:) | |
| 654 | apply (subst subst.simps) | |
| 655 | apply (subst match_Var_simps) | |
| 656 | apply (simp only: option.cases) | |
| 657 | apply (subst match_App_simps) | |
| 658 | apply (simp only: option.cases) | |
| 659 | apply (simp only: prod.cases) | |
| 660 | apply (simp only: lam.perm) | |
| 661 | apply (subst (3) subst.simps) | |
| 662 | apply (subst match_Var_simps) | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 663 | apply (simp only: option.cases) | 
| 2165 | 664 | apply (subst match_App_simps) | 
| 665 | apply (simp only: option.cases) | |
| 666 | apply (simp only: prod.cases) | |
| 667 | apply (subst 1(2)[of "(l, r)" "l" "r"]) | |
| 668 | apply (simp add: match_Var_simps) | |
| 669 | apply (simp add: match_App_simps) | |
| 670 | apply (rule refl) | |
| 671 | apply (subst 1(3)[of "(l, r)" "l" "r"]) | |
| 672 | apply (simp add: match_Var_simps) | |
| 673 | apply (simp add: match_App_simps) | |
| 674 | apply (rule refl) | |
| 675 | apply (rule refl) | |
| 676 | done | |
| 677 | next | |
| 678 | fix n t' | |
| 679 | assume "t = Lam n t'" | |
| 680 | then show ?thesis | |
| 681 | apply (simp only: ) | |
| 682 | apply (simp only: lam.perm) | |
| 683 | apply (subst subst.simps) | |
| 684 | apply (subst match_Var_simps) | |
| 685 | apply (simp only: option.cases) | |
| 686 | apply (subst match_App_simps) | |
| 687 | apply (simp only: option.cases) | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 688 | apply (rule_tac t="Lam n t'" and s="Lam (new ((v, s), Lam n t')) ((n \<leftrightarrow> new ((v, s), Lam n t')) \<bullet> t')" in subst) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 689 | defer | 
| 2165 | 690 | apply (subst match_Lam_simps) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 691 | defer | 
| 2165 | 692 | apply (simp only: option.cases) | 
| 693 | apply (simp only: prod.cases) | |
| 694 | apply (subst (2) subst.simps) | |
| 695 | apply (subst match_Var_simps) | |
| 696 | apply (simp only: option.cases) | |
| 697 | apply (subst match_App_simps) | |
| 698 | apply (simp only: option.cases) | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 699 | apply (rule_tac t="Lam (p \<bullet> n) (p \<bullet> t')" and s="Lam (new ((p \<bullet> v, p \<bullet> s), Lam (p \<bullet> n) (p \<bullet> t'))) (((p \<bullet> n) \<leftrightarrow> new ((p \<bullet> v, p \<bullet> s), Lam (p \<bullet> n) (p \<bullet> t'))) \<bullet> t')" in subst) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 700 | defer | 
| 2165 | 701 | apply (subst match_Lam_simps) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 702 | defer | 
| 2165 | 703 | apply (simp only: option.cases) | 
| 704 | apply (simp only: prod.cases) | |
| 705 | apply (simp only: lam.perm) | |
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 706 | thm 1(1) | 
| 2165 | 707 | sorry | 
| 708 | qed | |
| 709 | qed | |
| 710 | ||
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 711 | lemma subst_proper_eqs: | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 712 | "subst y s (Var x) = (if x = y then s else (Var x))" | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 713 | "subst y s (App l r) = App (subst y s l) (subst y s r)" | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 714 | "atom x \<sharp> (t, s) \<Longrightarrow> subst y s (Lam x t) = Lam x (subst y s t)" | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 715 | apply (subst subst.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 716 | apply (simp only: match_Var_simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 717 | apply (simp only: option.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 718 | apply (subst subst.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 719 | apply (simp only: match_App_simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 720 | apply (simp only: option.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 721 | apply (simp only: prod.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 722 | apply (simp only: match_Var_simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 723 | apply (simp only: option.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 724 | apply (subst subst.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 725 | apply (simp only: match_Var_simps) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 726 | apply (simp only: option.simps) | 
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 727 | apply (simp only: match_App_simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 728 | apply (simp only: option.simps) | 
| 2173 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 729 | apply (rule_tac t="Lam x t" and s="Lam (new ((y, s), Lam x t)) ((x \<leftrightarrow> new ((y, s), Lam x t)) \<bullet> t)" in subst) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 730 | defer | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 731 | apply (subst match_Lam_simps) | 
| 
477293d841e8
Match_Lam defined on Quotient Level.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2172diff
changeset | 732 | defer | 
| 2172 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 733 | apply (simp only: option.simps) | 
| 
fd5eec72c3f5
More on Function-defined subst.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
2170diff
changeset | 734 | apply (simp only: prod.simps) | 
| 2157 | 735 | sorry | 
| 2436 
3885dc2669f9
cleaned up (almost completely) the examples
 Christian Urban <urbanc@in.tum.de> parents: 
2434diff
changeset | 736 | *) | 
| 1594 | 737 | end | 
| 738 | ||
| 739 | ||
| 740 |