| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Thu, 26 Nov 2009 20:32:56 +0100 | |
| changeset 401 | 211229d6c66f | 
| parent 389 | d67240113f68 | 
| child 416 | 3f3927f793d4 | 
| permissions | -rw-r--r-- | 
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory LamEx | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | imports Nominal QuotMain | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | atom_decl name | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | |
| 243 | 7 | thm abs_fresh(1) | 
| 8 | ||
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | nominal_datatype rlam = | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | rVar "name" | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 | | rApp "rlam" "rlam" | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | | rLam "name" "rlam" | 
| 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | |
| 259 | 14 | print_theorems | 
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 | |
| 243 | 16 | function | 
| 17 | rfv :: "rlam \<Rightarrow> name set" | |
| 18 | where | |
| 19 |   rfv_var: "rfv (rVar a) = {a}"
 | |
| 20 | | rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)" | |
| 21 | | rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
 | |
| 22 | sorry | |
| 23 | ||
| 247 | 24 | termination rfv sorry | 
| 243 | 25 | |
| 271 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 26 | inductive | 
| 246 
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
 Christian Urban <urbanc@in.tum.de> parents: 
245diff
changeset | 27 |   alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
 | 
| 
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
 Christian Urban <urbanc@in.tum.de> parents: 
245diff
changeset | 28 | where | 
| 
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
 Christian Urban <urbanc@in.tum.de> parents: 
245diff
changeset | 29 | a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)" | 
| 
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
 Christian Urban <urbanc@in.tum.de> parents: 
245diff
changeset | 30 | | a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2" | 
| 
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
 Christian Urban <urbanc@in.tum.de> parents: 
245diff
changeset | 31 | | a3: "\<lbrakk>t \<approx> ([(a,b)]\<bullet>s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s" | 
| 
6da7d538e5e0
changed the order of rfv and reformulated a3 with rfv
 Christian Urban <urbanc@in.tum.de> parents: 
245diff
changeset | 32 | |
| 259 | 33 | print_theorems | 
| 34 | ||
| 271 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 35 | lemma alpha_refl: | 
| 286 | 36 | fixes t::"rlam" | 
| 37 | shows "t \<approx> t" | |
| 38 | apply(induct t rule: rlam.induct) | |
| 39 | apply(simp add: a1) | |
| 40 | apply(simp add: a2) | |
| 41 | apply(rule a3) | |
| 42 | apply(subst pt_swap_bij'') | |
| 43 | apply(rule pt_name_inst) | |
| 44 | apply(rule at_name_inst) | |
| 45 | apply(simp) | |
| 46 | apply(simp) | |
| 47 | done | |
| 48 | ||
| 49 | lemma alpha_EQUIV: | |
| 50 | shows "EQUIV alpha" | |
| 271 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 51 | sorry | 
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 52 | |
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 53 | quotient lam = rlam / alpha | 
| 286 | 54 | apply(rule alpha_EQUIV) | 
| 55 | done | |
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 56 | |
| 203 
7384115df9fd
added equiv-thm to the quot_info
 Christian Urban <urbanc@in.tum.de> parents: 
201diff
changeset | 57 | print_quotients | 
| 
7384115df9fd
added equiv-thm to the quot_info
 Christian Urban <urbanc@in.tum.de> parents: 
201diff
changeset | 58 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
267diff
changeset | 59 | quotient_def | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 60 | Var :: "name \<Rightarrow> lam" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 61 | where | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 62 | "Var \<equiv> rVar" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 63 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
267diff
changeset | 64 | quotient_def | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 65 | App :: "lam \<Rightarrow> lam \<Rightarrow> lam" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 66 | where | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 67 | "App \<equiv> rApp" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 68 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
267diff
changeset | 69 | quotient_def | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 70 | Lam :: "name \<Rightarrow> lam \<Rightarrow> lam" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 71 | where | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 72 | "Lam \<equiv> rLam" | 
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
203diff
changeset | 74 | thm Var_def | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
203diff
changeset | 75 | thm App_def | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
203diff
changeset | 76 | thm Lam_def | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
203diff
changeset | 77 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
267diff
changeset | 78 | quotient_def | 
| 243 | 79 | fv :: "lam \<Rightarrow> name set" | 
| 80 | where | |
| 81 | "fv \<equiv> rfv" | |
| 82 | ||
| 83 | thm fv_def | |
| 84 | ||
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 85 | (* definition of overloaded permutation function *) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 86 | (* for the lifted type lam *) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 87 | overloading | 
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
267diff
changeset | 88 | perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam" (unchecked) | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 89 | begin | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 90 | |
| 268 
4d58c02289ca
simplified the quotient_def code; type of the defined constant must now be given; for-part eliminated
 Christian Urban <urbanc@in.tum.de> parents: 
267diff
changeset | 91 | quotient_def | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 92 | perm_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 93 | where | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 94 | "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)" | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 95 | |
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 96 | end | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 97 | |
| 238 
e9cc3a3aa5d1
Tried manually lifting real_alpha
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
237diff
changeset | 98 | (*quotient_def (for lam) | 
| 
e9cc3a3aa5d1
Tried manually lifting real_alpha
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
237diff
changeset | 99 | abs_fun_lam :: "'x prm \<Rightarrow> lam \<Rightarrow> lam" | 
| 
e9cc3a3aa5d1
Tried manually lifting real_alpha
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
237diff
changeset | 100 | where | 
| 
e9cc3a3aa5d1
Tried manually lifting real_alpha
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
237diff
changeset | 101 | "perm_lam \<equiv> (perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam)"*) | 
| 
e9cc3a3aa5d1
Tried manually lifting real_alpha
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
237diff
changeset | 102 | |
| 
e9cc3a3aa5d1
Tried manually lifting real_alpha
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
237diff
changeset | 103 | |
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 104 | thm perm_lam_def | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 105 | |
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 106 | (* lemmas that need to lift *) | 
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 107 | lemma pi_var_com: | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 108 | fixes pi::"'x prm" | 
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 109 | shows "(pi\<bullet>rVar a) \<approx> rVar (pi\<bullet>a)" | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 110 | sorry | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 111 | |
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 112 | lemma pi_app_com: | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 113 | fixes pi::"'x prm" | 
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 114 | shows "(pi\<bullet>rApp t1 t2) \<approx> rApp (pi\<bullet>t1) (pi\<bullet>t2)" | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 115 | sorry | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 116 | |
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 117 | lemma pi_lam_com: | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 118 | fixes pi::"'x prm" | 
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 119 | shows "(pi\<bullet>rLam a t) \<approx> rLam (pi\<bullet>a) (pi\<bullet>t)" | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 120 | sorry | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 121 | |
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 122 | lemma real_alpha: | 
| 242 | 123 | assumes "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s" | 
| 201 
1ac36993cc71
added an example about lambda-terms
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 124 | shows "Lam a t = Lam b s" | 
| 217 
9cc87d02190a
First experiments with Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
203diff
changeset | 125 | sorry | 
| 
9cc87d02190a
First experiments with Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
203diff
changeset | 126 | |
| 286 | 127 | lemma perm_rsp: | 
| 128 | "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>" | |
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 129 | apply(auto) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 130 | (* this is propably true if some type conditions are imposed ;o) *) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 131 | sorry | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 132 | |
| 286 | 133 | lemma fresh_rsp: | 
| 134 | "(op = ===> alpha ===> op =) fresh fresh" | |
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 135 | apply(auto) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 136 | (* this is probably only true if some type conditions are imposed *) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 137 | sorry | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 138 | |
| 234 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 139 | lemma rVar_rsp: "(op = ===> alpha) rVar rVar" | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 140 | apply(auto) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 141 | apply(rule a1) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 142 | apply(simp) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 143 | done | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 144 | |
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 145 | lemma rApp_rsp: "(alpha ===> alpha ===> alpha) rApp rApp" | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 146 | apply(auto) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 147 | apply(rule a2) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 148 | apply (assumption) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 149 | apply (assumption) | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 150 | done | 
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 151 | |
| 
811f17de4031
Lifting of the 3 lemmas in LamEx
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
233diff
changeset | 152 | lemma rLam_rsp: "(op = ===> alpha ===> alpha) rLam rLam" | 
| 229 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 153 | apply(auto) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 154 | apply(rule a3) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 155 | apply(rule_tac t="[(x,x)]\<bullet>y" and s="y" in subst) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 156 | apply(rule sym) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 157 | apply(rule trans) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 158 | apply(rule pt_name3) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 159 | apply(rule at_ds1[OF at_name_inst]) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 160 | apply(simp add: pt_name1) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 161 | apply(assumption) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 162 | apply(simp add: abs_fresh) | 
| 
13f985a93dbc
fixed the definition of alpha; this *breaks* some of the experiments
 Christian Urban <urbanc@in.tum.de> parents: 
225diff
changeset | 163 | done | 
| 217 
9cc87d02190a
First experiments with Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
203diff
changeset | 164 | |
| 247 | 165 | lemma rfv_rsp: "(alpha ===> op =) rfv rfv" | 
| 166 | sorry | |
| 217 
9cc87d02190a
First experiments with Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
203diff
changeset | 167 | |
| 285 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 168 | lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)" | 
| 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 169 | apply (auto) | 
| 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 170 | apply (erule alpha.cases) | 
| 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 171 | apply (simp_all add: rlam.inject alpha_refl) | 
| 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 172 | done | 
| 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 173 | |
| 217 
9cc87d02190a
First experiments with Lambda
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
203diff
changeset | 174 | ML {* val qty = @{typ "lam"} *}
 | 
| 247 | 175 | ML {* val defs = @{thms Var_def App_def Lam_def perm_lam_def fv_def} *}
 | 
| 271 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 176 | ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} @
 | 
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 177 |   @{thms ho_all_prs ho_ex_prs} *}
 | 
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 178 | |
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 179 | ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *}
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 180 | ML {* val consts = lookup_quot_consts defs *}
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 181 | ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *}
 | 
| 389 | 182 | ML {* fun lift_tac_lam lthy t = lift_tac lthy t rel_eqv rel_refl rty quot trans2 rsp_thms reps_same absrep defs *}
 | 
| 237 | 183 | |
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 184 | lemma pi_var: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Var a = Var (pi \<bullet> a)"
 | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 185 | apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 186 | done | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 187 | |
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 188 | lemma pi_app: "(pi\<Colon>('x \<times> 'x) list) \<bullet> App (x\<Colon>lam) (xa\<Colon>lam) = App (pi \<bullet> x) (pi \<bullet> xa)"
 | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 189 | apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 190 | done | 
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 191 | |
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 192 | lemma pi_lam: "(pi\<Colon>('x \<times> 'x) list) \<bullet> Lam (a\<Colon>name) (x\<Colon>lam) = Lam (pi \<bullet> a) (pi \<bullet> x)"
 | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 193 | apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 194 | done | 
| 249 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 195 | |
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 196 | lemma fv_var: "fv (Var (a\<Colon>name)) = {a}"
 | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 197 | apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 198 | done | 
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 199 | |
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 200 | lemma fv_app: "fv (App (x\<Colon>lam) (xa\<Colon>lam)) = fv x \<union> fv xa" | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 201 | apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *})
 | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 202 | done | 
| 249 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 203 | |
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 204 | lemma fv_lam: "fv (Lam (a\<Colon>name) (x\<Colon>lam)) = fv x - {a}"
 | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 205 | apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *})
 | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 206 | done | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 207 | |
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 208 | lemma a1: "(a\<Colon>name) = (b\<Colon>name) \<Longrightarrow> Var a = Var b" | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 209 | apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 210 | done | 
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 211 | |
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 212 | lemma a2: "\<lbrakk>(x\<Colon>lam) = (xa\<Colon>lam); (xb\<Colon>lam) = (xc\<Colon>lam)\<rbrakk> \<Longrightarrow> App x xb = App xa xc" | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 213 | apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 214 | done | 
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 215 | |
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 216 | lemma a3: "\<lbrakk>(x\<Colon>lam) = [(a\<Colon>name, b\<Colon>name)] \<bullet> (xa\<Colon>lam); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa" | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 217 | apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *})
 | 
| 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 218 | done | 
| 286 | 219 | |
| 378 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 220 | lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P; | 
| 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 221 | \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P; | 
| 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 222 | \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk> | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 223 | \<Longrightarrow> P" | 
| 378 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 224 | apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *})
 | 
| 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 225 | done | 
| 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 226 | |
| 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 227 | lemma alpha_induct: "\<lbrakk>(qx\<Colon>lam) = (qxa\<Colon>lam); \<And>(a\<Colon>name) b\<Colon>name. a = b \<Longrightarrow> (qxb\<Colon>lam \<Rightarrow> lam \<Rightarrow> bool) (Var a) (Var b); | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 228 | \<And>(x\<Colon>lam) (xa\<Colon>lam) (xb\<Colon>lam) xc\<Colon>lam. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc); | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 229 | \<And>(x\<Colon>lam) (a\<Colon>name) (b\<Colon>name) xa\<Colon>lam. | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 230 | \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk> | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 231 | \<Longrightarrow> qxb qx qxa" | 
| 378 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 232 | apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *})
 | 
| 
86fba2c4eeef
All examples work again.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
376diff
changeset | 233 | done | 
| 252 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 234 | |
| 376 
e99c0334d8bf
lambda_prs and cleaning the existing examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
370diff
changeset | 235 | lemma var_inject: "(Var a = Var b) = (a = b)" | 
| 370 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 236 | apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *})
 | 
| 
09e28d4c19aa
Lambda & SOLVED' for new quotient_tac
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
292diff
changeset | 237 | done | 
| 285 
8ebdef196fd5
Infrastructure for polymorphic types
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
273diff
changeset | 238 | |
| 249 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 239 | lemma var_supp: | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 240 | shows "supp (Var a) = ((supp a)::name set)" | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 241 | apply(simp add: supp_def) | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 242 | apply(simp add: pi_var) | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 243 | apply(simp add: var_inject) | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 244 | done | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 245 | |
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 246 | lemma var_fresh: | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 247 | fixes a::"name" | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 248 | shows "(a\<sharp>(Var b)) = (a\<sharp>b)" | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 249 | apply(simp add: fresh_def) | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 250 | apply(simp add: var_supp) | 
| 
7dec34d12328
added some facts about fresh and support of lam
 Christian Urban <urbanc@in.tum.de> parents: 
247diff
changeset | 251 | done | 
| 247 | 252 | |
| 253 | ||
| 254 | ||
| 255 | ||
| 256 | ||
| 257 | ||
| 258 | ||
| 271 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 259 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 260 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 261 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 262 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 263 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 264 | (* Construction Site code *) | 
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 265 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 266 | |
| 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 267 | |
| 259 | 268 | |
| 269 | ||
| 270 | ||
| 271 | ||
| 271 
1b57f99737fe
Alpha.induct now lifts automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
270diff
changeset | 272 | |
| 237 | 273 | fun | 
| 274 |   option_map::"('a \<Rightarrow> 'b) \<Rightarrow> ('a noption) \<Rightarrow> ('b noption)"
 | |
| 275 | where | |
| 276 | "option_map f (nSome x) = nSome (f x)" | |
| 277 | | "option_map f nNone = nNone" | |
| 278 | ||
| 279 | fun | |
| 280 | option_rel | |
| 281 | where | |
| 282 | "option_rel r (nSome x) (nSome y) = r x y" | |
| 283 | | "option_rel r _ _ = False" | |
| 284 | ||
| 285 | declare [[map noption = (option_map, option_rel)]] | |
| 286 | ||
| 379 
57bde65f6eb2
Removed unused things from QuotMain.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
378diff
changeset | 287 | lemma "option_map id = id" | 
| 
57bde65f6eb2
Removed unused things from QuotMain.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
378diff
changeset | 288 | sorry | 
| 
57bde65f6eb2
Removed unused things from QuotMain.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
378diff
changeset | 289 | |
| 237 | 290 | lemma OPT_QUOTIENT: | 
| 291 | assumes q: "QUOTIENT R Abs Rep" | |
| 292 | shows "QUOTIENT (option_rel R) (option_map Abs) (option_map Rep)" | |
| 293 | apply (unfold QUOTIENT_def) | |
| 294 | apply (auto) | |
| 295 | using q | |
| 296 | apply (unfold QUOTIENT_def) | |
| 297 | apply (case_tac "a :: 'b noption") | |
| 298 | apply (simp) | |
| 299 | apply (simp) | |
| 300 | apply (case_tac "a :: 'b noption") | |
| 301 | apply (simp only: option_map.simps) | |
| 302 | apply (subst option_rel.simps) | |
| 303 | (* Simp starts hanging so don't know how to continue *) | |
| 304 | sorry | |
| 305 |