| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Fri, 26 Mar 2010 09:23:23 +0100 | |
| changeset 1652 | 3b9c05d158f3 | 
| parent 1593 | 20221ec06cba | 
| child 1656 | c9d3dda79fe3 | 
| permissions | -rw-r--r-- | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory LamEx | 
| 1652 
3b9c05d158f3
Fix Manual/LamEx for experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1593diff
changeset | 2 | imports "../Nominal2_Atoms" "../Nominal2_Eqvt" "../Nominal2_Supp" "../Abs" | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | atom_decl name | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | datatype rlam = | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | rVar "name" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | | rApp "rlam" "rlam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | | rLam "name" "rlam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | fun | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | rfv :: "rlam \<Rightarrow> atom set" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 14 | where | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 |   rfv_var: "rfv (rVar a) = {atom a}"
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | | rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | | rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}"
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 19 | instantiation rlam :: pt | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 20 | begin | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 21 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 22 | primrec | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 23 | permute_rlam | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 24 | where | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 25 | "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 26 | | "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 27 | | "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 28 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 29 | instance | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 30 | apply default | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 31 | apply(induct_tac [!] x) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 32 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 33 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 35 | end | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 36 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 37 | instantiation rlam :: fs | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 38 | begin | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | lemma neg_conj: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 41 | "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 42 | by simp | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 43 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 44 | instance | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 45 | apply default | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 46 | apply(induct_tac x) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 47 | (* var case *) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 48 | apply(simp add: supp_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 49 | apply(fold supp_def)[1] | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 50 | apply(simp add: supp_at_base) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 51 | (* app case *) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 52 | apply(simp only: supp_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | apply(simp only: permute_rlam.simps) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | apply(simp only: rlam.inject) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 55 | apply(simp only: neg_conj) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | apply(simp only: Collect_disj_eq) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 57 | apply(simp only: infinite_Un) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 58 | apply(simp only: Collect_disj_eq) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | apply(simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 60 | (* lam case *) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 61 | apply(simp only: supp_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 62 | apply(simp only: permute_rlam.simps) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | apply(simp only: rlam.inject) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | apply(simp only: neg_conj) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 65 | apply(simp only: Collect_disj_eq) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 66 | apply(simp only: infinite_Un) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 67 | apply(simp only: Collect_disj_eq) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 68 | apply(simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 69 | apply(fold supp_def)[1] | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 70 | apply(simp add: supp_at_base) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 71 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 72 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | end | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 75 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 76 | (* for the eqvt proof of the alpha-equivalence *) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 77 | declare permute_rlam.simps[eqvt] | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 78 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 79 | lemma rfv_eqvt[eqvt]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 80 | shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 81 | apply(induct t) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 82 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 83 | apply(simp add: permute_set_eq atom_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 84 | apply(simp add: union_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 85 | apply(simp add: Diff_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 86 | apply(simp add: permute_set_eq atom_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 87 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 88 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 89 | inductive | 
| 1652 
3b9c05d158f3
Fix Manual/LamEx for experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1593diff
changeset | 90 |     alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx>a _" [100, 100] 100)
 | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 91 | where | 
| 1652 
3b9c05d158f3
Fix Manual/LamEx for experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1593diff
changeset | 92 | a1: "a = b \<Longrightarrow> (rVar a) \<approx>a (rVar b)" | 
| 
3b9c05d158f3
Fix Manual/LamEx for experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1593diff
changeset | 93 | | a2: "\<lbrakk>t1 \<approx>a t2; s1 \<approx>a s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx>a rApp t2 s2" | 
| 
3b9c05d158f3
Fix Manual/LamEx for experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1593diff
changeset | 94 | | a3: "\<exists>pi. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx>a s)
 | 
| 
3b9c05d158f3
Fix Manual/LamEx for experiments.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1593diff
changeset | 95 | \<Longrightarrow> rLam a t \<approx>a rLam b s" | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 96 | |
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 97 | lemma a3_inverse: | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 98 | assumes "rLam a t \<approx> rLam b s" | 
| 1011 | 99 |   shows "\<exists>pi. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s)"
 | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 100 | using assms | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 101 | apply(erule_tac alpha.cases) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 102 | apply(auto) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 103 | done | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 104 | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 105 | text {* should be automatic with new version of eqvt-machinery *}
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 106 | lemma alpha_eqvt: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 107 | shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 108 | apply(induct rule: alpha.induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 109 | apply(simp add: a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 110 | apply(simp add: a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 111 | apply(simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 112 | apply(rule a3) | 
| 1011 | 113 | apply(erule conjE) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 114 | apply(erule exE) | 
| 1011 | 115 | apply(erule conjE) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 116 | apply(rule_tac x="pi \<bullet> pia" in exI) | 
| 1011 | 117 | apply(rule conjI) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 118 | apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1]) | 
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1028diff
changeset | 119 | apply(simp only: Diff_eqvt rfv_eqvt insert_eqvt atom_eqvt empty_eqvt) | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1028diff
changeset | 120 | apply(simp) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 121 | apply(rule conjI) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 122 | apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1]) | 
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1028diff
changeset | 123 | apply(simp add: Diff_eqvt rfv_eqvt atom_eqvt insert_eqvt empty_eqvt) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 124 | apply(subst permute_eqvt[symmetric]) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 125 | apply(simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 126 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 127 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 128 | lemma alpha_refl: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 129 | shows "t \<approx> t" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 130 | apply(induct t rule: rlam.induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 131 | apply(simp add: a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 132 | apply(simp add: a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 133 | apply(rule a3) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 134 | apply(rule_tac x="0" in exI) | 
| 1011 | 135 | apply(simp_all add: fresh_star_def fresh_zero_perm) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 136 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 137 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 138 | lemma alpha_sym: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 139 | shows "t \<approx> s \<Longrightarrow> s \<approx> t" | 
| 1011 | 140 | apply(induct rule: alpha.induct) | 
| 141 | apply(simp add: a1) | |
| 142 | apply(simp add: a2) | |
| 143 | apply(rule a3) | |
| 144 | apply(erule exE) | |
| 145 | apply(rule_tac x="- pi" in exI) | |
| 146 | apply(simp) | |
| 147 | apply(simp add: fresh_star_def fresh_minus_perm) | |
| 148 | apply(erule conjE)+ | |
| 149 | apply(rotate_tac 3) | |
| 150 | apply(drule_tac pi="- pi" in alpha_eqvt) | |
| 151 | apply(simp) | |
| 152 | done | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 153 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 154 | lemma alpha_trans: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 155 | shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 156 | apply(induct arbitrary: t3 rule: alpha.induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 157 | apply(erule alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 158 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 159 | apply(simp add: a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 160 | apply(rotate_tac 4) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 161 | apply(erule alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 162 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 163 | apply(simp add: a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 164 | apply(rotate_tac 1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 165 | apply(erule alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 166 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 167 | apply(erule conjE)+ | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 168 | apply(erule exE)+ | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 169 | apply(erule conjE)+ | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 170 | apply(rule a3) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 171 | apply(rule_tac x="pia + pi" in exI) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 172 | apply(simp add: fresh_star_plus) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 173 | apply(drule_tac x="- pia \<bullet> sa" in spec) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 174 | apply(drule mp) | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 175 | apply(rotate_tac 7) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 176 | apply(drule_tac pi="- pia" in alpha_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 177 | apply(simp) | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 178 | apply(rotate_tac 9) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 179 | apply(drule_tac pi="pia" in alpha_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 180 | apply(simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 181 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 182 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 183 | lemma alpha_equivp: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 184 | shows "equivp alpha" | 
| 981 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 185 | apply(rule equivpI) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 186 | unfolding reflp_def symp_def transp_def | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 187 | apply(auto intro: alpha_refl alpha_sym alpha_trans) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 188 | done | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 189 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 190 | lemma alpha_rfv: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 191 | shows "t \<approx> s \<Longrightarrow> rfv t = rfv s" | 
| 981 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 192 | apply(induct rule: alpha.induct) | 
| 1011 | 193 | apply(simp_all) | 
| 981 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 194 | done | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 195 | |
| 1011 | 196 | inductive | 
| 197 |     alpha2 :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx>2 _" [100, 100] 100)
 | |
| 198 | where | |
| 199 | a21: "a = b \<Longrightarrow> (rVar a) \<approx>2 (rVar b)" | |
| 200 | | a22: "\<lbrakk>t1 \<approx>2 t2; s1 \<approx>2 s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx>2 rApp t2 s2" | |
| 201 | | a23: "(a = b \<and> t \<approx>2 s) \<or> (a \<noteq> b \<and> ((a \<leftrightarrow> b) \<bullet> t) \<approx>2 s \<and> atom b \<notin> rfv t)\<Longrightarrow> rLam a t \<approx>2 rLam b s" | |
| 202 | ||
| 203 | lemma fv_vars: | |
| 204 | fixes a::name | |
| 205 |   assumes a1: "\<forall>x \<in> rfv t - {atom a}. pi \<bullet> x = x"
 | |
| 206 | shows "(pi \<bullet> t) \<approx>2 ((a \<leftrightarrow> (pi \<bullet> a)) \<bullet> t)" | |
| 207 | using a1 | |
| 208 | apply(induct t) | |
| 209 | apply(auto) | |
| 210 | apply(rule a21) | |
| 211 | apply(case_tac "name = a") | |
| 212 | apply(simp) | |
| 213 | apply(simp) | |
| 214 | defer | |
| 215 | apply(rule a22) | |
| 216 | apply(simp) | |
| 217 | apply(simp) | |
| 218 | apply(rule a23) | |
| 219 | apply(case_tac "a = name") | |
| 220 | apply(simp) | |
| 221 | oops | |
| 222 | ||
| 223 | ||
| 224 | lemma | |
| 225 | assumes a1: "t \<approx>2 s" | |
| 226 | shows "t \<approx> s" | |
| 227 | using a1 | |
| 228 | apply(induct) | |
| 229 | apply(rule alpha.intros) | |
| 230 | apply(simp) | |
| 231 | apply(rule alpha.intros) | |
| 232 | apply(simp) | |
| 233 | apply(simp) | |
| 234 | apply(rule alpha.intros) | |
| 235 | apply(erule disjE) | |
| 236 | apply(rule_tac x="0" in exI) | |
| 237 | apply(simp add: fresh_star_def fresh_zero_perm) | |
| 238 | apply(erule conjE)+ | |
| 239 | apply(drule alpha_rfv) | |
| 240 | apply(simp) | |
| 241 | apply(rule_tac x="(a \<leftrightarrow> b)" in exI) | |
| 242 | apply(simp) | |
| 243 | apply(erule conjE)+ | |
| 244 | apply(rule conjI) | |
| 245 | apply(drule alpha_rfv) | |
| 246 | apply(drule sym) | |
| 247 | apply(simp) | |
| 248 | apply(simp add: rfv_eqvt[symmetric]) | |
| 249 | defer | |
| 250 | apply(subgoal_tac "atom a \<sharp> (rfv t - {atom a})")
 | |
| 251 | apply(subgoal_tac "atom b \<sharp> (rfv t - {atom a})")
 | |
| 252 | ||
| 253 | defer | |
| 254 | sorry | |
| 255 | ||
| 256 | lemma | |
| 257 | assumes a1: "t \<approx> s" | |
| 258 | shows "t \<approx>2 s" | |
| 259 | using a1 | |
| 260 | apply(induct) | |
| 261 | apply(rule alpha2.intros) | |
| 262 | apply(simp) | |
| 263 | apply(rule alpha2.intros) | |
| 264 | apply(simp) | |
| 265 | apply(simp) | |
| 266 | apply(clarify) | |
| 267 | apply(rule alpha2.intros) | |
| 268 | apply(frule alpha_rfv) | |
| 269 | apply(rotate_tac 4) | |
| 270 | apply(drule sym) | |
| 271 | apply(simp) | |
| 272 | apply(drule sym) | |
| 273 | apply(simp) | |
| 274 | oops | |
| 275 | ||
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 276 | quotient_type lam = rlam / alpha | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 277 | by (rule alpha_equivp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 278 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 279 | quotient_definition | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 280 | "Var :: name \<Rightarrow> lam" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 281 | is | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 282 | "rVar" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 283 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 284 | quotient_definition | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 285 | "App :: lam \<Rightarrow> lam \<Rightarrow> lam" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 286 | is | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 287 | "rApp" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 288 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 289 | quotient_definition | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 290 | "Lam :: name \<Rightarrow> lam \<Rightarrow> lam" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 291 | is | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 292 | "rLam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 293 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 294 | quotient_definition | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 295 | "fv :: lam \<Rightarrow> atom set" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 296 | is | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 297 | "rfv" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 298 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 299 | lemma perm_rsp[quot_respect]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 300 | "(op = ===> alpha ===> alpha) permute permute" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 301 | apply(auto) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 302 | apply(rule alpha_eqvt) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 303 | apply(simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 304 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 305 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 306 | lemma rVar_rsp[quot_respect]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 307 | "(op = ===> alpha) rVar rVar" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 308 | by (auto intro: a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 309 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 310 | lemma rApp_rsp[quot_respect]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 311 | "(alpha ===> alpha ===> alpha) rApp rApp" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 312 | by (auto intro: a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 313 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 314 | lemma rLam_rsp[quot_respect]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 315 | "(op = ===> alpha ===> alpha) rLam rLam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 316 | apply(auto) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 317 | apply(rule a3) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 318 | apply(rule_tac x="0" in exI) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 319 | unfolding fresh_star_def | 
| 1011 | 320 | apply(simp add: fresh_star_def fresh_zero_perm) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 321 | apply(simp add: alpha_rfv) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 322 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 323 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 324 | lemma rfv_rsp[quot_respect]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 325 | "(alpha ===> op =) rfv rfv" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 326 | apply(simp add: alpha_rfv) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 327 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 328 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 329 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 330 | section {* lifted theorems *}
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 331 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 332 | lemma lam_induct: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 333 | "\<lbrakk>\<And>name. P (Var name); | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 334 | \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2); | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 335 | \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 336 | \<Longrightarrow> P lam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 337 | apply (lifting rlam.induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 338 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 339 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 340 | instantiation lam :: pt | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 341 | begin | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 342 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 343 | quotient_definition | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 344 | "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 345 | is | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 346 | "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 347 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 348 | lemma permute_lam [simp]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 349 | shows "pi \<bullet> Var a = Var (pi \<bullet> a)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 350 | and "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 351 | and "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 352 | apply(lifting permute_rlam.simps) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 353 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 354 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 355 | instance | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 356 | apply default | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 357 | apply(induct_tac [!] x rule: lam_induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 358 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 359 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 360 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 361 | end | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 362 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 363 | lemma fv_lam [simp]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 364 |   shows "fv (Var a) = {atom a}"
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 365 | and "fv (App t1 t2) = fv t1 \<union> fv t2" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 366 |   and   "fv (Lam a t) = fv t - {atom a}"
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 367 | apply(lifting rfv_var rfv_app rfv_lam) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 368 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 369 | |
| 990 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 370 | lemma fv_eqvt: | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 371 | shows "(p \<bullet> fv t) = fv (p \<bullet> t)" | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 372 | apply(lifting rfv_eqvt) | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 373 | done | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 374 | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 375 | lemma a1: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 376 | "a = b \<Longrightarrow> Var a = Var b" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 377 | by (lifting a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 378 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 379 | lemma a2: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 380 | "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 381 | by (lifting a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 382 | |
| 1011 | 383 | lemma a3: | 
| 384 |   "\<lbrakk>\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s)\<rbrakk> 
 | |
| 385 | \<Longrightarrow> Lam a t = Lam b s" | |
| 386 | apply (lifting a3) | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 387 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 388 | |
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 389 | lemma a3_inv: | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 390 | assumes "Lam a t = Lam b s" | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 391 |   shows "\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s)"
 | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 392 | using assms | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 393 | apply(lifting a3_inverse) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 394 | done | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 395 | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 396 | lemma alpha_cases: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 397 | "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P; | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 398 | \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P; | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 399 | \<And>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s; | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 400 |          \<exists>pi. fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a}) \<sharp>* pi \<and> (pi \<bullet> t) = s\<rbrakk> 
 | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 401 | \<Longrightarrow> P\<rbrakk> | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 402 | \<Longrightarrow> P" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 403 | by (lifting alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 404 | |
| 975 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 405 | (* not sure whether needed *) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 406 | lemma alpha_induct: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 407 | "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b); | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 408 | \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc); | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 409 | \<And>t a s b. | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 410 |         \<lbrakk>\<exists>pi. fv t - {atom a} = fv s - {atom b} \<and>
 | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 411 |          (fv t - {atom a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s)\<rbrakk> 
 | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 412 | \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk> | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 413 | \<Longrightarrow> qxb qx qxa" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 414 | by (lifting alpha.induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 415 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 416 | (* should they lift automatically *) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 417 | lemma lam_inject [simp]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 418 | shows "(Var a = Var b) = (a = b)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 419 | and "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 420 | apply(lifting rlam.inject(1) rlam.inject(2)) | 
| 981 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 421 | apply(regularize) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 422 | prefer 2 | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 423 | apply(regularize) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 424 | prefer 2 | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 425 | apply(auto) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 426 | apply(drule alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 427 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 428 | apply(simp add: alpha.a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 429 | apply(drule alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 430 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 431 | apply(drule alpha.cases) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 432 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 433 | apply(rule alpha.a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 434 | apply(simp_all) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 435 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 436 | |
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 437 | lemma Lam_pseudo_inject: | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 438 | shows "(Lam a t = Lam b s) = | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 439 |       (\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s))"
 | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 440 | apply(rule iffI) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 441 | apply(rule a3_inv) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 442 | apply(assumption) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 443 | apply(rule a3) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 444 | apply(assumption) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 445 | done | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 446 | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 447 | lemma rlam_distinct: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 448 | shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 449 | and "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 450 | and "\<not>(rVar nam \<approx> rLam nam' rlam')" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 451 | and "\<not>(rLam nam' rlam' \<approx> rVar nam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 452 | and "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 453 | and "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 454 | apply auto | 
| 981 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 455 | apply (erule alpha.cases) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 456 | apply (simp_all only: rlam.distinct) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 457 | apply (erule alpha.cases) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 458 | apply (simp_all only: rlam.distinct) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 459 | apply (erule alpha.cases) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 460 | apply (simp_all only: rlam.distinct) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 461 | apply (erule alpha.cases) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 462 | apply (simp_all only: rlam.distinct) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 463 | apply (erule alpha.cases) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 464 | apply (simp_all only: rlam.distinct) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 465 | apply (erule alpha.cases) | 
| 
bc739536b715
Minor when looking at lam.distinct and lam.inject
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
975diff
changeset | 466 | apply (simp_all only: rlam.distinct) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 467 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 468 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 469 | lemma lam_distinct[simp]: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 470 | shows "Var nam \<noteq> App lam1' lam2'" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 471 | and "App lam1' lam2' \<noteq> Var nam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 472 | and "Var nam \<noteq> Lam nam' lam'" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 473 | and "Lam nam' lam' \<noteq> Var nam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 474 | and "App lam1 lam2 \<noteq> Lam nam' lam'" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 475 | and "Lam nam' lam' \<noteq> App lam1 lam2" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 476 | apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6)) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 477 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 478 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 479 | lemma var_supp1: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 480 | shows "(supp (Var a)) = (supp a)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 481 | apply (simp add: supp_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 482 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 483 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 484 | lemma var_supp: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 485 |   shows "(supp (Var a)) = {a:::name}"
 | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 486 | using var_supp1 by (simp add: supp_at_base) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 487 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 488 | lemma app_supp: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 489 | shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)" | 
| 997 
b7d259ded92e
Ported LF to the generic lambda and solved the simpler _supp cases.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
995diff
changeset | 490 | apply(simp only: supp_def lam_inject) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 491 | apply(simp add: Collect_imp_eq Collect_neg_eq) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 492 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 493 | |
| 990 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 494 | (* supp for lam *) | 
| 975 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 495 | lemma lam_supp1: | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 496 | shows "(supp (atom x, t)) supports (Lam x t) " | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 497 | apply(simp add: supports_def) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 498 | apply(fold fresh_def) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 499 | apply(simp add: fresh_Pair swap_fresh_fresh) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 500 | apply(clarify) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 501 | apply(subst swap_at_base_simps(3)) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 502 | apply(simp_all add: fresh_atom) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 503 | done | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 504 | |
| 975 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 505 | lemma lam_fsupp1: | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 506 | assumes a: "finite (supp t)" | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 507 | shows "finite (supp (Lam x t))" | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 508 | apply(rule supports_finite) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 509 | apply(rule lam_supp1) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 510 | apply(simp add: a supp_Pair supp_atom) | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 511 | done | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 512 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 513 | instance lam :: fs | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 514 | apply(default) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 515 | apply(induct_tac x rule: lam_induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 516 | apply(simp add: var_supp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 517 | apply(simp add: app_supp) | 
| 975 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 518 | apply(simp add: lam_fsupp1) | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 519 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 520 | |
| 990 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 521 | lemma supp_fv: | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 522 | shows "supp t = fv t" | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 523 | apply(induct t rule: lam_induct) | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 524 | apply(simp add: var_supp) | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 525 | apply(simp add: app_supp) | 
| 1016 
de8da5b32265
repaired according to changes in Abs.thy
 Christian Urban <urbanc@in.tum.de> parents: 
1011diff
changeset | 526 | apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)")
 | 
| 
de8da5b32265
repaired according to changes in Abs.thy
 Christian Urban <urbanc@in.tum.de> parents: 
1011diff
changeset | 527 | apply(simp add: supp_Abs) | 
| 990 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 528 | apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt) | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 529 | apply(simp add: Lam_pseudo_inject) | 
| 1016 
de8da5b32265
repaired according to changes in Abs.thy
 Christian Urban <urbanc@in.tum.de> parents: 
1011diff
changeset | 530 | apply(simp add: Abs_eq_iff alpha_gen) | 
| 990 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 531 | apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric]) | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 532 | done | 
| 
c25ff084868f
now also final step is proved - the supp of lambdas is now completely characterised
 Christian Urban <urbanc@in.tum.de> parents: 
989diff
changeset | 533 | |
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 534 | lemma lam_supp2: | 
| 1016 
de8da5b32265
repaired according to changes in Abs.thy
 Christian Urban <urbanc@in.tum.de> parents: 
1011diff
changeset | 535 |   shows "supp (Lam x t) = supp (Abs {atom x} t)"
 | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 536 | apply(simp add: supp_def permute_set_eq atom_eqvt) | 
| 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 537 | apply(simp add: Lam_pseudo_inject) | 
| 1016 
de8da5b32265
repaired according to changes in Abs.thy
 Christian Urban <urbanc@in.tum.de> parents: 
1011diff
changeset | 538 | apply(simp add: Abs_eq_iff supp_fv alpha_gen) | 
| 995 
ee0619b5adff
introduced a generic alpha (but not sure whether it is helpful)
 Christian Urban <urbanc@in.tum.de> parents: 
990diff
changeset | 539 | done | 
| 975 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 540 | |
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 541 | lemma lam_supp: | 
| 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 542 |   shows "supp (Lam x t) = ((supp t) - {atom x})"
 | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 543 | apply(simp add: lam_supp2) | 
| 1016 
de8da5b32265
repaired according to changes in Abs.thy
 Christian Urban <urbanc@in.tum.de> parents: 
1011diff
changeset | 544 | apply(simp add: supp_Abs) | 
| 989 
af02b193a19a
the supp of a lambda can now be characterised, *provided* the notion of free variables coincides with support on lambda terms
 Christian Urban <urbanc@in.tum.de> parents: 
984diff
changeset | 545 | done | 
| 975 
513ebe332964
test about supp/freshness for lam (old proofs work in principle - for single binders)
 Christian Urban <urbanc@in.tum.de> parents: 
947diff
changeset | 546 | |
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 547 | lemma fresh_lam: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 548 | "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 549 | apply(simp add: fresh_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 550 | apply(simp add: lam_supp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 551 | apply(auto) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 552 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 553 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 554 | lemma lam_induct_strong: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 555 | fixes a::"'a::fs" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 556 | assumes a1: "\<And>name b. P b (Var name)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 557 | and a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 558 | and a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 559 | shows "P a lam" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 560 | proof - | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 561 | have "\<And>pi a. P a (pi \<bullet> lam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 562 | proof (induct lam rule: lam_induct) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 563 | case (1 name pi) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 564 | show "P a (pi \<bullet> Var name)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 565 | apply (simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 566 | apply (rule a1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 567 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 568 | next | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 569 | case (2 lam1 lam2 pi) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 570 | have b1: "\<And>pi a. P a (pi \<bullet> lam1)" by fact | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 571 | have b2: "\<And>pi a. P a (pi \<bullet> lam2)" by fact | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 572 | show "P a (pi \<bullet> App lam1 lam2)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 573 | apply (simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 574 | apply (rule a2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 575 | apply (rule b1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 576 | apply (rule b2) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 577 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 578 | next | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 579 | case (3 name lam pi a) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 580 | have b: "\<And>pi a. P a (pi \<bullet> lam)" by fact | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 581 | obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 582 | apply(rule obtain_atom) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 583 | apply(auto) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 584 | sorry | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 585 | from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 586 | apply - | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 587 | apply(rule a3) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 588 | apply(blast) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 589 | apply(simp add: fresh_Pair) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 590 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 591 | have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 592 | apply(rule swap_fresh_fresh) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 593 | using fr | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 594 | apply(simp add: fresh_lam fresh_Pair) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 595 | apply(simp add: fresh_lam fresh_Pair) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 596 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 597 | show "P a (pi \<bullet> Lam name lam)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 598 | apply (simp) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 599 | apply(subst eq[symmetric]) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 600 | using p | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 601 | apply(simp only: permute_lam) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 602 | apply(simp add: flip_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 603 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 604 | qed | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 605 | then have "P a (0 \<bullet> lam)" by blast | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 606 | then show "P a lam" by simp | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 607 | qed | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 608 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 609 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 610 | lemma var_fresh: | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 611 | fixes a::"name" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 612 | shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)" | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 613 | apply(simp add: fresh_def) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 614 | apply(simp add: var_supp1) | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 615 | done | 
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 616 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 617 | |
| 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 618 | |
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1028diff
changeset | 619 | end | 
| 947 
fa810f01f7b5
added an LamEx example together with the new nominal infrastructure
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 620 |