| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Wed, 17 Feb 2010 16:22:16 +0100 | |
| changeset 1180 | 3f36936f1280 | 
| parent 1179 | 789fbba5c23f | 
| child 1181 | 788a59d2d7aa | 
| permissions | -rw-r--r-- | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory Terms | 
| 1180 | 2 | imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | atom_decl name | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | text {* primrec seems to be genarally faster than fun *}
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | section {*** lets with binding patterns ***}
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | |
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 11 | datatype rtrm1 = | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 12 | rVr1 "name" | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 13 | | rAp1 "rtrm1" "rtrm1" | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 14 | | rLm1 "name" "rtrm1" --"name is bound in trm1" | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 15 | | rLt1 "bp" "rtrm1" "rtrm1" --"all variables in bp are bound in the 2nd trm1" | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | and bp = | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | BUnit | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | | BVr "name" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 19 | | BPr "bp" "bp" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 20 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 21 | (* to be given by the user *) | 
| 1029 | 22 | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 23 | primrec | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 24 | bv1 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 25 | where | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 26 |   "bv1 (BUnit) = {}"
 | 
| 957 
080bd6f1607c
mostly ported Terms.thy to new Nominal
 Christian Urban <urbanc@in.tum.de> parents: 
950diff
changeset | 27 | | "bv1 (BVr x) = {atom x}"
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 28 | | "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 29 | |
| 1180 | 30 | local_setup {* define_raw_fv "Terms.rtrm1"
 | 
| 31 |   [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(NONE, 0)], [], [(SOME @{term bv1}, 0)]]],
 | |
| 32 | [[], [[]], [[], []]]] *} | |
| 33 | print_theorems | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 35 | setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Terms.rtrm1", "Terms.bp"] *}
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 36 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 37 | inductive | 
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 38 |   alpha1 :: "rtrm1 \<Rightarrow> rtrm1 \<Rightarrow> bool" ("_ \<approx>1 _" [100, 100] 100)
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 | where | 
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 40 | a1: "a = b \<Longrightarrow> (rVr1 a) \<approx>1 (rVr1 b)" | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 41 | | a2: "\<lbrakk>t1 \<approx>1 t2; s1 \<approx>1 s2\<rbrakk> \<Longrightarrow> rAp1 t1 s1 \<approx>1 rAp1 t2 s2" | 
| 1180 | 42 | | a3: "(\<exists>pi. (({atom aa}, t) \<approx>gen alpha1 fv_rtrm1 pi ({atom ab}, s))) \<Longrightarrow> rLm1 aa t \<approx>1 rLm1 ab s"
 | 
| 43 | | a4: "t1 \<approx>1 t2 \<Longrightarrow> (\<exists>pi. (((bv1 b1), s1) \<approx>gen alpha1 fv_rtrm1 pi ((bv1 b2), s2))) \<Longrightarrow> rLt1 b1 t1 s1 \<approx>1 rLt1 b2 t2 s2" | |
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 44 | |
| 1029 | 45 | lemma alpha1_inj: | 
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 46 | "(rVr1 a \<approx>1 rVr1 b) = (a = b)" | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 47 | "(rAp1 t1 s1 \<approx>1 rAp1 t2 s2) = (t1 \<approx>1 t2 \<and> s1 \<approx>1 s2)" | 
| 1180 | 48 | "(rLm1 aa t \<approx>1 rLm1 ab s) = (\<exists>pi. (({atom aa}, t) \<approx>gen alpha1 fv_rtrm1 pi ({atom ab}, s)))"
 | 
| 49 | "(rLt1 b1 t1 s1 \<approx>1 rLt1 b2 t2 s2) = (t1 \<approx>1 t2 \<and> (\<exists>pi. (((bv1 b1), s1) \<approx>gen alpha1 fv_rtrm1 pi ((bv1 b2), s2))))" | |
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 50 | apply - | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 51 | apply rule apply (erule alpha1.cases) apply (simp_all add: alpha1.intros) | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 52 | apply rule apply (erule alpha1.cases) apply (simp_all add: alpha1.intros) | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 53 | apply rule apply (erule alpha1.cases) apply (simp_all add: alpha1.intros) | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 54 | apply rule apply (erule alpha1.cases) apply (simp_all add: alpha1.intros) | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 55 | done | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 56 | |
| 1030 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 57 | (* Shouyld we derive it? But bv is given by the user? *) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 58 | lemma bv1_eqvt[eqvt]: | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 59 | shows "(pi \<bullet> bv1 x) = bv1 (pi \<bullet> x)" | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 60 | apply (induct x) | 
| 1043 
534d4c604f80
fixed proofs that broke because of eqvt
 Christian Urban <urbanc@in.tum.de> parents: 
1039diff
changeset | 61 | apply (simp_all add: empty_eqvt insert_eqvt atom_eqvt) | 
| 1030 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 62 | done | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 63 | |
| 1180 | 64 | lemma fv_rtrm1_eqvt[eqvt]: | 
| 65 | shows "(pi\<bullet>fv_rtrm1 t) = fv_rtrm1 (pi\<bullet>t)" | |
| 1031 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 66 | apply (induct t) | 
| 1043 
534d4c604f80
fixed proofs that broke because of eqvt
 Christian Urban <urbanc@in.tum.de> parents: 
1039diff
changeset | 67 | apply (simp_all add: insert_eqvt atom_eqvt empty_eqvt union_eqvt Diff_eqvt bv1_eqvt) | 
| 1031 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 68 | done | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 69 | |
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 70 | |
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 71 | lemma alpha1_eqvt: | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 72 | shows "t \<approx>1 s \<Longrightarrow> (pi \<bullet> t) \<approx>1 (pi \<bullet> s)" | 
| 1031 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 73 | apply (induct t s rule: alpha1.inducts) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 74 | apply (simp_all add:eqvts alpha1_inj) | 
| 1033 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 75 | apply (erule exE) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 76 | apply (rule_tac x="pi \<bullet> pia" in exI) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 77 | apply (simp add: alpha_gen) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 78 | apply(erule conjE)+ | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 79 | apply(rule conjI) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 80 | apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1]) | 
| 1180 | 81 | apply(simp add: atom_eqvt Diff_eqvt insert_eqvt empty_eqvt fv_rtrm1_eqvt) | 
| 1033 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 82 | apply(rule conjI) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 83 | apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1]) | 
| 1180 | 84 | apply(simp add: atom_eqvt Diff_eqvt fv_rtrm1_eqvt insert_eqvt empty_eqvt) | 
| 1033 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 85 | apply(simp add: permute_eqvt[symmetric]) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 86 | apply (erule exE) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 87 | apply (rule_tac x="pi \<bullet> pia" in exI) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 88 | apply (simp add: alpha_gen) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 89 | apply(erule conjE)+ | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 90 | apply(rule conjI) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 91 | apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1]) | 
| 1180 | 92 | apply(simp add: fv_rtrm1_eqvt Diff_eqvt bv1_eqvt) | 
| 1033 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 93 | apply(rule conjI) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 94 | apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1]) | 
| 1180 | 95 | apply(simp add: atom_eqvt fv_rtrm1_eqvt Diff_eqvt bv1_eqvt) | 
| 1033 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 96 | apply(simp add: permute_eqvt[symmetric]) | 
| 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 97 | done | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 98 | |
| 957 
080bd6f1607c
mostly ported Terms.thy to new Nominal
 Christian Urban <urbanc@in.tum.de> parents: 
950diff
changeset | 99 | lemma alpha1_equivp: "equivp alpha1" | 
| 
080bd6f1607c
mostly ported Terms.thy to new Nominal
 Christian Urban <urbanc@in.tum.de> parents: 
950diff
changeset | 100 | sorry | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 101 | |
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 102 | quotient_type trm1 = rtrm1 / alpha1 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 103 | by (rule alpha1_equivp) | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 104 | |
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 105 | quotient_definition | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 106 | "Vr1 :: name \<Rightarrow> trm1" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 107 | is | 
| 1028 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 108 | "rVr1" | 
| 
41fc4d3fc764
First experiments in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
976diff
changeset | 109 | |
| 1029 | 110 | quotient_definition | 
| 111 | "Ap1 :: trm1 \<Rightarrow> trm1 \<Rightarrow> trm1" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 112 | is | 
| 1029 | 113 | "rAp1" | 
| 114 | ||
| 115 | quotient_definition | |
| 116 | "Lm1 :: name \<Rightarrow> trm1 \<Rightarrow> trm1" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 117 | is | 
| 1029 | 118 | "rLm1" | 
| 119 | ||
| 120 | quotient_definition | |
| 121 | "Lt1 :: bp \<Rightarrow> trm1 \<Rightarrow> trm1 \<Rightarrow> trm1" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 122 | is | 
| 1029 | 123 | "rLt1" | 
| 124 | ||
| 125 | quotient_definition | |
| 126 | "fv_trm1 :: trm1 \<Rightarrow> atom set" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 127 | is | 
| 1180 | 128 | "fv_rtrm1" | 
| 1029 | 129 | |
| 130 | lemma alpha_rfv1: | |
| 1180 | 131 | shows "t \<approx>1 s \<Longrightarrow> fv_rtrm1 t = fv_rtrm1 s" | 
| 1029 | 132 | apply(induct rule: alpha1.induct) | 
| 133 | apply(simp_all add: alpha_gen.simps) | |
| 134 | done | |
| 135 | ||
| 136 | lemma [quot_respect]: | |
| 137 | "(op = ===> alpha1) rVr1 rVr1" | |
| 138 | "(alpha1 ===> alpha1 ===> alpha1) rAp1 rAp1" | |
| 139 | "(op = ===> alpha1 ===> alpha1) rLm1 rLm1" | |
| 140 | "(op = ===> alpha1 ===> alpha1 ===> alpha1) rLt1 rLt1" | |
| 141 | apply (auto intro: alpha1.intros) | |
| 142 | apply(rule a3) apply (rule_tac x="0" in exI) | |
| 143 | apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv1 alpha_gen) | |
| 144 | apply(rule a4) apply assumption apply (rule_tac x="0" in exI) | |
| 145 | apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv1 alpha_gen) | |
| 146 | done | |
| 147 | ||
| 148 | lemma [quot_respect]: | |
| 149 | "(op = ===> alpha1 ===> alpha1) permute permute" | |
| 150 | apply auto | |
| 151 | apply (rule alpha1_eqvt) | |
| 152 | apply simp | |
| 153 | done | |
| 154 | ||
| 155 | lemma [quot_respect]: | |
| 1180 | 156 | "(alpha1 ===> op =) fv_rtrm1 fv_rtrm1" | 
| 1029 | 157 | apply (simp add: alpha_rfv1) | 
| 158 | done | |
| 159 | ||
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 160 | lemmas trm1_bp_induct = rtrm1_bp.induct[quot_lifted] | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 161 | lemmas trm1_bp_inducts = rtrm1_bp.inducts[quot_lifted] | 
| 1029 | 162 | |
| 163 | instantiation trm1 and bp :: pt | |
| 164 | begin | |
| 165 | ||
| 166 | quotient_definition | |
| 167 | "permute_trm1 :: perm \<Rightarrow> trm1 \<Rightarrow> trm1" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 168 | is | 
| 1029 | 169 | "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1" | 
| 170 | ||
| 1083 
30550327651a
Proper context fixes lifting inside instantiations.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1073diff
changeset | 171 | lemmas permute_trm1[simp] = permute_rtrm1_permute_bp.simps[quot_lifted] | 
| 
30550327651a
Proper context fixes lifting inside instantiations.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1073diff
changeset | 172 | |
| 1029 | 173 | instance | 
| 174 | apply default | |
| 175 | apply(induct_tac [!] x rule: trm1_bp_inducts(1)) | |
| 1083 
30550327651a
Proper context fixes lifting inside instantiations.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1073diff
changeset | 176 | apply(simp_all) | 
| 1029 | 177 | done | 
| 178 | ||
| 179 | end | |
| 180 | ||
| 1180 | 181 | lemmas fv_trm1 = fv_rtrm1_fv_bp.simps[quot_lifted] | 
| 1029 | 182 | |
| 1180 | 183 | lemmas fv_trm1_eqvt = fv_rtrm1_eqvt[quot_lifted] | 
| 1029 | 184 | |
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 185 | lemmas alpha1_INJ = alpha1_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] | 
| 1029 | 186 | |
| 1031 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 187 | lemma lm1_supp_pre: | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 188 | shows "(supp (atom x, t)) supports (Lm1 x t) " | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 189 | apply(simp add: supports_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 190 | apply(fold fresh_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 191 | apply(simp add: fresh_Pair swap_fresh_fresh) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 192 | apply(clarify) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 193 | apply(subst swap_at_base_simps(3)) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 194 | apply(simp_all add: fresh_atom) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 195 | done | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 196 | |
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 197 | lemma lt1_supp_pre: | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 198 | shows "(supp (x, t, s)) supports (Lt1 t x s) " | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 199 | apply(simp add: supports_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 200 | apply(fold fresh_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 201 | apply(simp add: fresh_Pair swap_fresh_fresh) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 202 | done | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 203 | |
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 204 | lemma bp_supp: "finite (supp (bp :: bp))" | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 205 | apply (induct bp) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 206 | apply(simp_all add: supp_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 207 | apply (fold supp_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 208 | apply (simp add: supp_at_base) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 209 | apply(simp add: Collect_imp_eq) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 210 | apply(simp add: Collect_neg_eq[symmetric]) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 211 | apply (fold supp_def) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 212 | apply (simp) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 213 | done | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 214 | |
| 1030 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 215 | instance trm1 :: fs | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 216 | apply default | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 217 | apply(induct_tac x rule: trm1_bp_inducts(1)) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 218 | apply(simp_all) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 219 | apply(simp add: supp_def alpha1_INJ eqvts) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 220 | apply(simp add: supp_def[symmetric] supp_at_base) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 221 | apply(simp only: supp_def alpha1_INJ eqvts permute_trm1) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 222 | apply(simp add: Collect_imp_eq Collect_neg_eq) | 
| 1031 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 223 | apply(rule supports_finite) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 224 | apply(rule lm1_supp_pre) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 225 | apply(simp add: supp_Pair supp_atom) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 226 | apply(rule supports_finite) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 227 | apply(rule lt1_supp_pre) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 228 | apply(simp add: supp_Pair supp_atom bp_supp) | 
| 
bcf6e7d20c20
More ingredients in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1030diff
changeset | 229 | done | 
| 1030 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 230 | |
| 1029 | 231 | lemma supp_fv: | 
| 232 | shows "supp t = fv_trm1 t" | |
| 233 | apply(induct t rule: trm1_bp_inducts(1)) | |
| 1030 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 234 | apply(simp_all) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 235 | apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 236 | apply(simp only: supp_at_base[simplified supp_def]) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 237 | apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 238 | apply(simp add: Collect_imp_eq Collect_neg_eq) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 239 | apply(subgoal_tac "supp (Lm1 name rtrm1) = supp (Abs {atom name} rtrm1)")
 | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 240 | apply(simp add: supp_Abs fv_trm1) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 241 | apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 242 | apply(simp add: alpha1_INJ) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 243 | apply(simp add: Abs_eq_iff) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 244 | apply(simp add: alpha_gen.simps) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 245 | apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric]) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 246 | apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp(rtrm11) \<union> supp (Abs (bv1 bp) rtrm12)") | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 247 | apply(simp add: supp_Abs fv_trm1) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 248 | apply(simp (no_asm) add: supp_def) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 249 | apply(simp add: alpha1_INJ) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 250 | apply(simp add: Abs_eq_iff) | 
| 1033 
dce45db16063
Some cleaning and eqvt proof
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1032diff
changeset | 251 | apply(simp add: alpha_gen) | 
| 1030 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 252 | apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 253 | apply(simp add: Collect_imp_eq Collect_neg_eq) | 
| 
07f97267a392
Finished the supp_fv proof; first proof that analyses the structure of 'Let' :)
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1029diff
changeset | 254 | done | 
| 1029 | 255 | |
| 1032 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 256 | lemma trm1_supp: | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 257 |   "supp (Vr1 x) = {atom x}"
 | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 258 | "supp (Ap1 t1 t2) = supp t1 \<union> supp t2" | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 259 |   "supp (Lm1 x t) = (supp t) - {atom x}"
 | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 260 | "supp (Lt1 b t s) = supp t \<union> (supp s - bv1 b)" | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 261 | by (simp_all only: supp_fv fv_trm1) | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 262 | |
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 263 | lemma trm1_induct_strong: | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 264 | assumes "\<And>name b. P b (Vr1 name)" | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 265 | and "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)" | 
| 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 266 | and "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)" | 
| 1179 
789fbba5c23f
Fix the strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1171diff
changeset | 267 | and "\<And>bp rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12; bv1 bp \<sharp>* b\<rbrakk> \<Longrightarrow> P b (Lt1 bp rtrm11 rtrm12)" | 
| 1032 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 268 | shows "P a rtrma" | 
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 269 | sorry | 
| 1032 
135bf399c036
The trm1_support lemma explicitly and stated a strong induction principle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1031diff
changeset | 270 | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 271 | section {*** lets with single assignments ***}
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 272 | |
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 273 | datatype rtrm2 = | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 274 | rVr2 "name" | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 275 | | rAp2 "rtrm2" "rtrm2" | 
| 1093 
139b257e10d2
Explicitly marked what is bound.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1092diff
changeset | 276 | | rLm2 "name" "rtrm2" --"bind (name) in (rtrm2)" | 
| 
139b257e10d2
Explicitly marked what is bound.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1092diff
changeset | 277 | | rLt2 "rassign" "rtrm2" --"bind (bv2 rassign) in (rtrm2)" | 
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 278 | and rassign = | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 279 | rAs "name" "rtrm2" | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 280 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 281 | (* to be given by the user *) | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 282 | primrec | 
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 283 | rbv2 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 284 | where | 
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 285 |   "rbv2 (rAs x t) = {atom x}"
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 286 | |
| 1180 | 287 | local_setup {* define_raw_fv "Terms.rtrm2"
 | 
| 288 |   [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv2}, 0)]]],
 | |
| 289 | [[[(NONE, 0)], []]]] *} | |
| 290 | print_theorems | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 291 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 292 | setup {* snd o define_raw_perms ["rtrm2", "rassign"] ["Terms.rtrm2", "Terms.rassign"] *}
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 293 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 294 | inductive | 
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 295 |   alpha2 :: "rtrm2 \<Rightarrow> rtrm2 \<Rightarrow> bool" ("_ \<approx>2 _" [100, 100] 100)
 | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 296 | and | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 297 |   alpha2a :: "rassign \<Rightarrow> rassign \<Rightarrow> bool" ("_ \<approx>2a _" [100, 100] 100)
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 298 | where | 
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 299 | a1: "a = b \<Longrightarrow> (rVr2 a) \<approx>2 (rVr2 b)" | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 300 | | a2: "\<lbrakk>t1 \<approx>2 t2; s1 \<approx>2 s2\<rbrakk> \<Longrightarrow> rAp2 t1 s1 \<approx>2 rAp2 t2 s2" | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 301 | | a3: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha2 fv_rtrm2 pi ({atom b}, s))) \<Longrightarrow> rLm2 a t \<approx>2 rLm2 b s"
 | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 302 | | a4: "\<lbrakk>\<exists>pi. ((rbv2 bt, t) \<approx>gen alpha2 fv_rtrm2 pi ((rbv2 bs), s)); | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 303 | \<exists>pi. ((rbv2 bt, bt) \<approx>gen alpha2a fv_rassign pi (rbv2 bs, bs))\<rbrakk> | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 304 | \<Longrightarrow> rLt2 bt t \<approx>2 rLt2 bs s" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 305 | | a5: "\<lbrakk>a = b; t \<approx>2 s\<rbrakk> \<Longrightarrow> rAs a t \<approx>2a rAs b s" (* This way rbv2 can be lifted *) | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 306 | |
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 307 | lemma alpha2_equivp: | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 308 | "equivp alpha2" | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 309 | "equivp alpha2a" | 
| 957 
080bd6f1607c
mostly ported Terms.thy to new Nominal
 Christian Urban <urbanc@in.tum.de> parents: 
950diff
changeset | 310 | sorry | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 311 | |
| 1091 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 312 | quotient_type | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 313 | trm2 = rtrm2 / alpha2 | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 314 | and | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 315 | assign = rassign / alpha2a | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 316 | by (auto intro: alpha2_equivp) | 
| 
d3946f1a9341
Looking at the trm2 example
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1083diff
changeset | 317 | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 318 | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 319 | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 320 | section {*** lets with many assignments ***}
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 321 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 322 | datatype trm3 = | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 323 | Vr3 "name" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 324 | | Ap3 "trm3" "trm3" | 
| 1093 
139b257e10d2
Explicitly marked what is bound.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1092diff
changeset | 325 | | Lm3 "name" "trm3" --"bind (name) in (trm3)" | 
| 
139b257e10d2
Explicitly marked what is bound.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1092diff
changeset | 326 | | Lt3 "assigns" "trm3" --"bind (bv3 assigns) in (trm3)" | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 327 | and assigns = | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 328 | ANil | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 329 | | ACons "name" "trm3" "assigns" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 330 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 331 | (* to be given by the user *) | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 332 | primrec | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 333 | bv3 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 334 | where | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 335 |   "bv3 ANil = {}"
 | 
| 957 
080bd6f1607c
mostly ported Terms.thy to new Nominal
 Christian Urban <urbanc@in.tum.de> parents: 
950diff
changeset | 336 | | "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)"
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 337 | |
| 1180 | 338 | local_setup {* define_raw_fv "Terms.trm3"
 | 
| 339 |   [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term bv3}, 0)]]],
 | |
| 340 | [[], [[(NONE, 0)], [], []]]] *} | |
| 341 | print_theorems | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 342 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 343 | setup {* snd o define_raw_perms ["rtrm3", "assigns"] ["Terms.trm3", "Terms.assigns"] *}
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 344 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 345 | inductive | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 346 |   alpha3 :: "trm3 \<Rightarrow> trm3 \<Rightarrow> bool" ("_ \<approx>3 _" [100, 100] 100)
 | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 347 | and | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 348 |   alpha3a :: "assigns \<Rightarrow> assigns \<Rightarrow> bool" ("_ \<approx>3a _" [100, 100] 100)
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 349 | where | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 350 | a1: "a = b \<Longrightarrow> (Vr3 a) \<approx>3 (Vr3 b)" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 351 | | a2: "\<lbrakk>t1 \<approx>3 t2; s1 \<approx>3 s2\<rbrakk> \<Longrightarrow> Ap3 t1 s1 \<approx>3 Ap3 t2 s2" | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 352 | | a3: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha3 fv_rtrm3 pi ({atom b}, s))) \<Longrightarrow> Lm3 a t \<approx>3 Lm3 b s"
 | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 353 | | a4: "\<lbrakk>\<exists>pi. ((bv3 bt, t) \<approx>gen alpha3 fv_trm3 pi ((bv3 bs), s)); | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 354 | \<exists>pi. ((bv3 bt, bt) \<approx>gen alpha3a fv_assign pi (bv3 bs, bs))\<rbrakk> | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 355 | \<Longrightarrow> Lt3 bt t \<approx>3 Lt3 bs s" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 356 | | a5: "ANil \<approx>3a ANil" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 357 | | a6: "\<lbrakk>a = b; t \<approx>3 s; tt \<approx>3a st\<rbrakk> \<Longrightarrow> ACons a t tt \<approx>3a ACons b s st" | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 358 | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 359 | lemma alpha3_equivp: | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 360 | "equivp alpha3" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 361 | "equivp alpha3a" | 
| 957 
080bd6f1607c
mostly ported Terms.thy to new Nominal
 Christian Urban <urbanc@in.tum.de> parents: 
950diff
changeset | 362 | sorry | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 363 | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 364 | quotient_type | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 365 | qtrm3 = trm3 / alpha3 | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 366 | and | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 367 | qassigns = assigns / alpha3a | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 368 | by (auto intro: alpha3_equivp) | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 369 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 370 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 371 | section {*** lam with indirect list recursion ***}
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 372 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 373 | datatype trm4 = | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 374 | Vr4 "name" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 375 | | Ap4 "trm4" "trm4 list" | 
| 1093 
139b257e10d2
Explicitly marked what is bound.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1092diff
changeset | 376 | | Lm4 "name" "trm4" --"bind (name) in (trm)" | 
| 1180 | 377 | print_theorems | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 378 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 379 | thm trm4.recs | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 380 | |
| 1180 | 381 | local_setup {* define_raw_fv "Terms.trm4" [
 | 
| 382 | [[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]]], [[], [[], []]] ] *} | |
| 383 | print_theorems | |
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 384 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 385 | (* there cannot be a clause for lists, as *) | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 386 | (* permutations are already defined in Nominal (also functions, options, and so on) *) | 
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 387 | setup {* snd o define_raw_perms ["trm4"] ["Terms.trm4"] *}
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 388 | |
| 963 | 389 | (* "repairing" of the permute function *) | 
| 390 | lemma repaired: | |
| 391 | fixes ts::"trm4 list" | |
| 392 | shows "permute_trm4_list p ts = p \<bullet> ts" | |
| 393 | apply(induct ts) | |
| 394 | apply(simp_all) | |
| 395 | done | |
| 396 | ||
| 976 | 397 | thm permute_trm4_permute_trm4_list.simps | 
| 963 | 398 | thm permute_trm4_permute_trm4_list.simps[simplified repaired] | 
| 399 | ||
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 400 | inductive | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 401 |     alpha4 :: "trm4 \<Rightarrow> trm4 \<Rightarrow> bool" ("_ \<approx>4 _" [100, 100] 100)
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 402 | and alpha4list :: "trm4 list \<Rightarrow> trm4 list \<Rightarrow> bool" ("_ \<approx>4list _" [100, 100] 100) 
 | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 403 | where | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 404 | a1: "a = b \<Longrightarrow> (Vr4 a) \<approx>4 (Vr4 b)" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 405 | | a2: "\<lbrakk>t1 \<approx>4 t2; s1 \<approx>4list s2\<rbrakk> \<Longrightarrow> Ap4 t1 s1 \<approx>4 Ap4 t2 s2" | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 406 | | a3: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha4 fv_rtrm4 pi ({atom b}, s))) \<Longrightarrow> Lm4 a t \<approx>4 Lm4 b s"
 | 
| 950 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 407 | | a5: "[] \<approx>4list []" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 408 | | a6: "\<lbrakk>t \<approx>4 s; ts \<approx>4list ss\<rbrakk> \<Longrightarrow> (t#ts) \<approx>4list (s#ss)" | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 409 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 410 | lemma alpha4_equivp: "equivp alpha4" sorry | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 411 | lemma alpha4list_equivp: "equivp alpha4list" sorry | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 412 | |
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 413 | quotient_type | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 414 | qtrm4 = trm4 / alpha4 and | 
| 
98764f25f012
added Terms to Nominal - Instantiation of two types does not work (ask Florian)
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 415 | qtrm4list = "trm4 list" / alpha4list | 
| 1042 | 416 | by (simp_all add: alpha4_equivp alpha4list_equivp) | 
| 1036 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 417 | |
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 418 | |
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 419 | datatype rtrm5 = | 
| 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 420 | rVr5 "name" | 
| 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 421 | | rAp5 "rtrm5" "rtrm5" | 
| 1093 
139b257e10d2
Explicitly marked what is bound.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1092diff
changeset | 422 | | rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)" | 
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 423 | and rlts = | 
| 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 424 | rLnil | 
| 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 425 | | rLcons "name" "rtrm5" "rlts" | 
| 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 426 | |
| 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 427 | primrec | 
| 1050 | 428 | rbv5 | 
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 429 | where | 
| 1050 | 430 |   "rbv5 rLnil = {}"
 | 
| 431 | | "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)"
 | |
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 432 | |
| 1180 | 433 | local_setup {* define_raw_fv "Terms.rtrm5" [
 | 
| 434 |   [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[(NONE, 0)], [], []]]  ] *}
 | |
| 435 | print_theorems | |
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 436 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 437 | setup {* snd o define_raw_perms ["rtrm5", "rlts"] ["Terms.rtrm5", "Terms.rlts"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 438 | print_theorems | 
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 439 | |
| 1036 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 440 | inductive | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 441 |   alpha5 :: "rtrm5 \<Rightarrow> rtrm5 \<Rightarrow> bool" ("_ \<approx>5 _" [100, 100] 100)
 | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 442 | and | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 443 |   alphalts :: "rlts \<Rightarrow> rlts \<Rightarrow> bool" ("_ \<approx>l _" [100, 100] 100)
 | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 444 | where | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 445 | a1: "a = b \<Longrightarrow> (rVr5 a) \<approx>5 (rVr5 b)" | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 446 | | a2: "\<lbrakk>t1 \<approx>5 t2; s1 \<approx>5 s2\<rbrakk> \<Longrightarrow> rAp5 t1 s1 \<approx>5 rAp5 t2 s2" | 
| 1180 | 447 | | a3: "\<lbrakk>\<exists>pi. ((rbv5 l1, t1) \<approx>gen alpha5 fv_rtrm5 pi (rbv5 l2, t2)); | 
| 448 | \<exists>pi. ((rbv5 l1, l1) \<approx>gen alphalts fv_rlts pi (rbv5 l2, l2))\<rbrakk> | |
| 1036 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 449 | \<Longrightarrow> rLt5 l1 t1 \<approx>5 rLt5 l2 t2" | 
| 
aaac8274f08c
The alpha-equivalence relation for let-rec. Not sure if correct...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1035diff
changeset | 450 | | a4: "rLnil \<approx>l rLnil" | 
| 1040 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 451 | | a5: "ls1 \<approx>l ls2 \<Longrightarrow> t1 \<approx>5 t2 \<Longrightarrow> n1 = n2 \<Longrightarrow> rLcons n1 t1 ls1 \<approx>l rLcons n2 t2 ls2" | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 452 | |
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 453 | print_theorems | 
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 454 | |
| 1040 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 455 | lemma alpha5_inj: | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 456 | "((rVr5 a) \<approx>5 (rVr5 b)) = (a = b)" | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 457 | "(rAp5 t1 s1 \<approx>5 rAp5 t2 s2) = (t1 \<approx>5 t2 \<and> s1 \<approx>5 s2)" | 
| 1180 | 458 | "(rLt5 l1 t1 \<approx>5 rLt5 l2 t2) = ((\<exists>pi. ((rbv5 l1, t1) \<approx>gen alpha5 fv_rtrm5 pi (rbv5 l2, t2))) \<and> | 
| 459 | (\<exists>pi. ((rbv5 l1, l1) \<approx>gen alphalts fv_rlts pi (rbv5 l2, l2))))" | |
| 1040 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 460 | "rLnil \<approx>l rLnil" | 
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 461 | "(rLcons n1 t1 ls1 \<approx>l rLcons n2 t2 ls2) = (n1 = n2 \<and> ls1 \<approx>l ls2 \<and> t1 \<approx>5 t2)" | 
| 1040 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 462 | apply - | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 463 | apply (simp_all add: alpha5_alphalts.intros) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 464 | apply rule | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 465 | apply (erule alpha5.cases) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 466 | apply (simp_all add: alpha5_alphalts.intros) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 467 | apply rule | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 468 | apply (erule alpha5.cases) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 469 | apply (simp_all add: alpha5_alphalts.intros) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 470 | apply rule | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 471 | apply (erule alpha5.cases) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 472 | apply (simp_all add: alpha5_alphalts.intros) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 473 | apply rule | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 474 | apply (erule alphalts.cases) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 475 | apply (simp_all add: alpha5_alphalts.intros) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 476 | done | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 477 | |
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 478 | lemma alpha5_equivps: | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 479 | shows "equivp alpha5" | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 480 | and "equivp alphalts" | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 481 | sorry | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 482 | |
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 483 | quotient_type | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 484 | trm5 = rtrm5 / alpha5 | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 485 | and | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 486 | lts = rlts / alphalts | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 487 | by (auto intro: alpha5_equivps) | 
| 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 488 | |
| 1046 | 489 | quotient_definition | 
| 490 | "Vr5 :: name \<Rightarrow> trm5" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 491 | is | 
| 1046 | 492 | "rVr5" | 
| 493 | ||
| 494 | quotient_definition | |
| 495 | "Ap5 :: trm5 \<Rightarrow> trm5 \<Rightarrow> trm5" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 496 | is | 
| 1046 | 497 | "rAp5" | 
| 498 | ||
| 499 | quotient_definition | |
| 500 | "Lt5 :: lts \<Rightarrow> trm5 \<Rightarrow> trm5" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 501 | is | 
| 1046 | 502 | "rLt5" | 
| 503 | ||
| 504 | quotient_definition | |
| 505 | "Lnil :: lts" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 506 | is | 
| 1046 | 507 | "rLnil" | 
| 508 | ||
| 509 | quotient_definition | |
| 510 | "Lcons :: name \<Rightarrow> trm5 \<Rightarrow> lts \<Rightarrow> lts" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 511 | is | 
| 1046 | 512 | "rLcons" | 
| 513 | ||
| 514 | quotient_definition | |
| 515 | "fv_trm5 :: trm5 \<Rightarrow> atom set" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 516 | is | 
| 1180 | 517 | "fv_rtrm5" | 
| 1046 | 518 | |
| 519 | quotient_definition | |
| 1050 | 520 | "fv_lts :: lts \<Rightarrow> atom set" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 521 | is | 
| 1180 | 522 | "fv_rlts" | 
| 1046 | 523 | |
| 1050 | 524 | quotient_definition | 
| 525 | "bv5 :: lts \<Rightarrow> atom set" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 526 | is | 
| 1050 | 527 | "rbv5" | 
| 528 | ||
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 529 | lemma rbv5_eqvt: | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 530 | "pi \<bullet> (rbv5 x) = rbv5 (pi \<bullet> x)" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 531 | sorry | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 532 | |
| 1180 | 533 | lemma fv_rtrm5_eqvt: | 
| 534 | "pi \<bullet> (fv_rtrm5 x) = fv_rtrm5 (pi \<bullet> x)" | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 535 | sorry | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 536 | |
| 1180 | 537 | lemma fv_rlts_eqvt: | 
| 538 | "pi \<bullet> (fv_rlts x) = fv_rlts (pi \<bullet> x)" | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 539 | sorry | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 540 | |
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 541 | lemma alpha5_eqvt: | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 542 | "xa \<approx>5 y \<Longrightarrow> (x \<bullet> xa) \<approx>5 (x \<bullet> y)" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 543 | "xb \<approx>l ya \<Longrightarrow> (x \<bullet> xb) \<approx>l (x \<bullet> ya)" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 544 | apply(induct rule: alpha5_alphalts.inducts) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 545 | apply (simp_all add: alpha5_inj) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 546 | apply (erule exE)+ | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 547 | apply(unfold alpha_gen) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 548 | apply (erule conjE)+ | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 549 | apply (rule conjI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 550 | apply (rule_tac x="x \<bullet> pi" in exI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 551 | apply (rule conjI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 552 | apply(rule_tac ?p1="- x" in permute_eq_iff[THEN iffD1]) | 
| 1180 | 553 | apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rtrm5_eqvt) | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 554 | apply(rule conjI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 555 | apply(rule_tac ?p1="- x" in fresh_star_permute_iff[THEN iffD1]) | 
| 1180 | 556 | apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rtrm5_eqvt) | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 557 | apply (subst permute_eqvt[symmetric]) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 558 | apply (simp) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 559 | apply (rule_tac x="x \<bullet> pia" in exI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 560 | apply (rule conjI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 561 | apply(rule_tac ?p1="- x" in permute_eq_iff[THEN iffD1]) | 
| 1180 | 562 | apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rlts_eqvt) | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 563 | apply(rule conjI) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 564 | apply(rule_tac ?p1="- x" in fresh_star_permute_iff[THEN iffD1]) | 
| 1180 | 565 | apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rlts_eqvt) | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 566 | apply (subst permute_eqvt[symmetric]) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 567 | apply (simp) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 568 | done | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 569 | |
| 1050 | 570 | lemma alpha5_rfv: | 
| 1180 | 571 | "(t \<approx>5 s \<Longrightarrow> fv_rtrm5 t = fv_rtrm5 s)" | 
| 572 | "(l \<approx>l m \<Longrightarrow> fv_rlts l = fv_rlts m)" | |
| 1050 | 573 | apply(induct rule: alpha5_alphalts.inducts) | 
| 574 | apply(simp_all add: alpha_gen) | |
| 575 | done | |
| 576 | ||
| 1053 
b1ca92ea3a86
proved that bv for lists respects alpha for terms
 Christian Urban <urbanc@in.tum.de> parents: 
1051diff
changeset | 577 | lemma bv_list_rsp: | 
| 1055 
34220518cccf
Simplified the proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1054diff
changeset | 578 | shows "x \<approx>l y \<Longrightarrow> rbv5 x = rbv5 y" | 
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 579 | apply(induct rule: alpha5_alphalts.inducts(2)) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 580 | apply(simp_all) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 581 | done | 
| 1053 
b1ca92ea3a86
proved that bv for lists respects alpha for terms
 Christian Urban <urbanc@in.tum.de> parents: 
1051diff
changeset | 582 | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 583 | lemma [quot_respect]: | 
| 1180 | 584 | "(alphalts ===> op =) fv_rlts fv_rlts" | 
| 585 | "(alpha5 ===> op =) fv_rtrm5 fv_rtrm5" | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 586 | "(alphalts ===> op =) rbv5 rbv5" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 587 | "(op = ===> alpha5) rVr5 rVr5" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 588 | "(alpha5 ===> alpha5 ===> alpha5) rAp5 rAp5" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 589 | "(alphalts ===> alpha5 ===> alpha5) rLt5 rLt5" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 590 | "(alphalts ===> alpha5 ===> alpha5) rLt5 rLt5" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 591 | "(op = ===> alpha5 ===> alphalts ===> alphalts) rLcons rLcons" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 592 | "(op = ===> alpha5 ===> alpha5) permute permute" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 593 | "(op = ===> alphalts ===> alphalts) permute permute" | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 594 | apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 595 | apply (auto) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 596 | apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 597 | apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 598 | apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 599 | apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv) | 
| 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 600 | done | 
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 601 | |
| 1092 
01ae4a87c7c3
Cleaning and updating in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1091diff
changeset | 602 | lemma | 
| 1053 
b1ca92ea3a86
proved that bv for lists respects alpha for terms
 Christian Urban <urbanc@in.tum.de> parents: 
1051diff
changeset | 603 | shows "(alphalts ===> op =) rbv5 rbv5" | 
| 
b1ca92ea3a86
proved that bv for lists respects alpha for terms
 Christian Urban <urbanc@in.tum.de> parents: 
1051diff
changeset | 604 | by (simp add: bv_list_rsp) | 
| 
b1ca92ea3a86
proved that bv for lists respects alpha for terms
 Christian Urban <urbanc@in.tum.de> parents: 
1051diff
changeset | 605 | |
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 606 | lemmas trm5_lts_inducts = rtrm5_rlts.inducts[quot_lifted] | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 607 | |
| 1050 | 608 | instantiation trm5 and lts :: pt | 
| 609 | begin | |
| 610 | ||
| 611 | quotient_definition | |
| 612 | "permute_trm5 :: perm \<Rightarrow> trm5 \<Rightarrow> trm5" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 613 | is | 
| 1050 | 614 | "permute :: perm \<Rightarrow> rtrm5 \<Rightarrow> rtrm5" | 
| 615 | ||
| 616 | quotient_definition | |
| 617 | "permute_lts :: perm \<Rightarrow> lts \<Rightarrow> lts" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 618 | is | 
| 1050 | 619 | "permute :: perm \<Rightarrow> rlts \<Rightarrow> rlts" | 
| 620 | ||
| 621 | lemma trm5_lts_zero: | |
| 622 | "0 \<bullet> (x\<Colon>trm5) = x" | |
| 623 | "0 \<bullet> (y\<Colon>lts) = y" | |
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 624 | apply(induct x and y rule: trm5_lts_inducts) | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 625 | apply(simp_all add: permute_rtrm5_permute_rlts.simps[quot_lifted]) | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 626 | done | 
| 1050 | 627 | |
| 628 | lemma trm5_lts_plus: | |
| 629 | "(p + q) \<bullet> (x\<Colon>trm5) = p \<bullet> q \<bullet> x" | |
| 630 | "(p + q) \<bullet> (y\<Colon>lts) = p \<bullet> q \<bullet> y" | |
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 631 | apply(induct x and y rule: trm5_lts_inducts) | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 632 | apply(simp_all add: permute_rtrm5_permute_rlts.simps[quot_lifted]) | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 633 | done | 
| 1050 | 634 | |
| 635 | instance | |
| 636 | apply default | |
| 637 | apply (simp_all add: trm5_lts_zero trm5_lts_plus) | |
| 638 | done | |
| 1040 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 639 | |
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 640 | end | 
| 1050 | 641 | |
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 642 | lemmas permute_trm5_lts = permute_rtrm5_permute_rlts.simps[quot_lifted] | 
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 643 | |
| 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 644 | lemmas alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] | 
| 1050 | 645 | |
| 1073 
53350d409473
Cleaned Terms using [lifted] and found a workaround for the instantiation problem.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1058diff
changeset | 646 | lemmas bv5[simp] = rbv5.simps[quot_lifted] | 
| 1050 | 647 | |
| 1180 | 648 | lemmas fv_trm5_lts[simp] = fv_rtrm5_fv_rlts.simps[quot_lifted] | 
| 1050 | 649 | |
| 650 | lemma lets_ok: | |
| 651 | "(Lt5 (Lcons x (Vr5 x) Lnil) (Vr5 x)) = (Lt5 (Lcons y (Vr5 y) Lnil) (Vr5 y))" | |
| 652 | apply (subst alpha5_INJ) | |
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 653 | apply (rule conjI) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 654 | apply (rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 655 | apply (simp only: alpha_gen) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 656 | apply (simp add: permute_trm5_lts fresh_star_def) | 
| 1050 | 657 | apply (rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 658 | apply (simp only: alpha_gen) | |
| 1052 
c1b469325033
Finished remains on the let proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1051diff
changeset | 659 | apply (simp add: permute_trm5_lts fresh_star_def) | 
| 
c1b469325033
Finished remains on the let proof.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1051diff
changeset | 660 | done | 
| 1050 | 661 | |
| 1058 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 662 | lemma lets_ok2: | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 663 | "(Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) = | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 664 | (Lt5 (Lcons y (Vr5 y) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 665 | apply (subst alpha5_INJ) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 666 | apply (rule conjI) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 667 | apply (rule_tac x="0 :: perm" in exI) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 668 | apply (simp only: alpha_gen) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 669 | apply (simp add: permute_trm5_lts fresh_star_def) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 670 | apply (rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 671 | apply (simp only: alpha_gen) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 672 | apply (simp add: permute_trm5_lts fresh_star_def) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 673 | done | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 674 | |
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 675 | |
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 676 | lemma lets_not_ok1: | 
| 1056 | 677 | "x \<noteq> y \<Longrightarrow> (Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq> | 
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 678 | (Lt5 (Lcons y (Vr5 x) (Lcons x (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 679 | apply (subst alpha5_INJ(3)) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 680 | apply(clarify) | 
| 1056 | 681 | apply (simp add: alpha_gen) | 
| 682 | apply (simp add: permute_trm5_lts fresh_star_def) | |
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 683 | apply (simp add: alpha5_INJ(5)) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 684 | apply(clarify) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 685 | apply (simp add: alpha5_INJ(2)) | 
| 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 686 | apply (simp only: alpha5_INJ(1)) | 
| 1056 | 687 | done | 
| 688 | ||
| 1058 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 689 | lemma distinct_helper: | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 690 | shows "\<not>(rVr5 x \<approx>5 rAp5 y z)" | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 691 | apply auto | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 692 | apply (erule alpha5.cases) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 693 | apply (simp_all only: rtrm5.distinct) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 694 | done | 
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 695 | |
| 1058 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 696 | lemma distinct_helper2: | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 697 | shows "(Vr5 x) \<noteq> (Ap5 y z)" | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 698 | by (lifting distinct_helper) | 
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 699 | |
| 1058 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 700 | lemma lets_nok: | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 701 | "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow> | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 702 | (Lt5 (Lcons x (Ap5 (Vr5 z) (Vr5 z)) (Lcons y (Vr5 z) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq> | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 703 | (Lt5 (Lcons y (Vr5 z) (Lcons x (Ap5 (Vr5 z) (Vr5 z)) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 704 | apply (subst alpha5_INJ) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 705 | apply (simp only: alpha_gen permute_trm5_lts fresh_star_def) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 706 | apply (subst alpha5_INJ(5)) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 707 | apply (subst alpha5_INJ(5)) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 708 | apply (simp add: distinct_helper2) | 
| 
afedef46d3ab
More let-rec experiments
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1057diff
changeset | 709 | done | 
| 1057 
f81b506f62a7
proposal for an alpha equivalence
 Christian Urban <urbanc@in.tum.de> parents: 
1056diff
changeset | 710 | |
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 711 | |
| 1105 | 712 | (* example with a bn function defined over the type itself *) | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 713 | datatype rtrm6 = | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 714 | rVr6 "name" | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 715 | | rLm6 "name" "rtrm6" | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 716 | | rLt6 "rtrm6" "rtrm6" --"bind (bv6 left) in (right)" | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 717 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 718 | primrec | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 719 | rbv6 | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 720 | where | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 721 |   "rbv6 (rVr6 n) = {}"
 | 
| 1117 
8948319b190e
Fixed rbv6, when translating to OTT.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1111diff
changeset | 722 | | "rbv6 (rLm6 n t) = {atom n} \<union> rbv6 t"
 | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 723 | | "rbv6 (rLt6 l r) = rbv6 l \<union> rbv6 r" | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 724 | |
| 1180 | 725 | local_setup {* define_raw_fv "Terms.rtrm6" [
 | 
| 726 |   [[[]], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv6}, 0)]]]] *}
 | |
| 727 | print_theorems | |
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 728 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 729 | setup {* snd o define_raw_perms ["rtrm6"] ["Terms.rtrm6"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 730 | print_theorems | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 731 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 732 | inductive | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 733 |   alpha6 :: "rtrm6 \<Rightarrow> rtrm6 \<Rightarrow> bool" ("_ \<approx>6 _" [100, 100] 100)
 | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 734 | where | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 735 | a1: "a = b \<Longrightarrow> (rVr6 a) \<approx>6 (rVr6 b)" | 
| 1180 | 736 | | a2: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha6 fv_rtrm6 pi ({atom b}, s))) \<Longrightarrow> rLm6 a t \<approx>6 rLm6 b s"
 | 
| 737 | | a3: "(\<exists>pi. (((rbv6 t1), s1) \<approx>gen alpha6 fv_rtrm6 pi ((rbv6 t2), s2))) \<Longrightarrow> rLt6 t1 s1 \<approx>6 rLt6 t2 s2" | |
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 738 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 739 | lemma alpha6_equivps: | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 740 | shows "equivp alpha6" | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 741 | sorry | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 742 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 743 | quotient_type | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 744 | trm6 = rtrm6 / alpha6 | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 745 | by (auto intro: alpha6_equivps) | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 746 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 747 | quotient_definition | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 748 | "Vr6 :: name \<Rightarrow> trm6" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 749 | is | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 750 | "rVr6" | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 751 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 752 | quotient_definition | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 753 | "Lm6 :: name \<Rightarrow> trm6 \<Rightarrow> trm6" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 754 | is | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 755 | "rLm6" | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 756 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 757 | quotient_definition | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 758 | "Lt6 :: trm6 \<Rightarrow> trm6 \<Rightarrow> trm6" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 759 | is | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 760 | "rLt6" | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 761 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 762 | quotient_definition | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 763 | "fv_trm6 :: trm6 \<Rightarrow> atom set" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 764 | is | 
| 1180 | 765 | "fv_rtrm6" | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 766 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 767 | quotient_definition | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 768 | "bv6 :: trm6 \<Rightarrow> atom set" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 769 | is | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 770 | "rbv6" | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 771 | |
| 1106 | 772 | lemma [quot_respect]: | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 773 | "(op = ===> alpha6 ===> alpha6) permute permute" | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 774 | apply auto (* will work with eqvt *) | 
| 1106 | 775 | sorry | 
| 776 | ||
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 777 | (* Definitely not true , see lemma below *) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 778 | |
| 1104 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 779 | lemma [quot_respect]:"(alpha6 ===> op =) rbv6 rbv6" | 
| 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 780 | apply simp apply clarify | 
| 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 781 | apply (erule alpha6.induct) | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 782 | oops | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 783 | |
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 784 | lemma "(a :: name) \<noteq> b \<Longrightarrow> \<not> (alpha6 ===> op =) rbv6 rbv6" | 
| 1104 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 785 | apply simp | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 786 | apply (rule_tac x="rLm6 (a::name) (rVr6 (a :: name))" in exI) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 787 | apply (rule_tac x="rLm6 (b::name) (rVr6 (b :: name))" in exI) | 
| 1104 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 788 | apply simp | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 789 | apply (rule a2) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 790 | apply (rule_tac x="(a \<leftrightarrow> b)" in exI) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 791 | apply (simp add: alpha_gen fresh_star_def) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 792 | apply (rule a1) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 793 | apply (rule refl) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 794 | done | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 795 | |
| 1180 | 796 | lemma [quot_respect]:"(alpha6 ===> op =) fv_rtrm6 fv_rtrm6" | 
| 1106 | 797 | apply simp apply clarify | 
| 798 | apply (induct_tac x y rule: alpha6.induct) | |
| 799 | apply simp_all | |
| 800 | apply (erule exE) | |
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 801 | apply (simp_all add: alpha_gen) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 802 | apply (erule conjE)+ | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 803 | apply (erule exE) | 
| 1106 | 804 | apply (erule conjE)+ | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 805 | apply (simp) | 
| 1106 | 806 | oops | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 807 | |
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 808 | |
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 809 | lemma [quot_respect]: "(op = ===> alpha6) rVr6 rVr6" | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 810 | by (simp_all add: a1) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 811 | |
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 812 | lemma [quot_respect]: | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 813 | "(op = ===> alpha6 ===> alpha6) rLm6 rLm6" | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 814 | "(alpha6 ===> alpha6 ===> alpha6) rLt6 rLt6" | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 815 | apply simp_all apply (clarify) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 816 | apply (rule a2) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 817 | apply (rule_tac x="0::perm" in exI) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 818 | apply (simp add: alpha_gen) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 819 | (* needs rfv6_rsp *) defer | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 820 | apply clarify | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 821 | apply (rule a3) | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 822 | apply (rule_tac x="0::perm" in exI) | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 823 | apply (simp add: alpha_gen) | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 824 | (* needs rbv6_rsp *) | 
| 1106 | 825 | oops | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 826 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 827 | instantiation trm6 :: pt begin | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 828 | |
| 1104 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 829 | quotient_definition | 
| 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 830 | "permute_trm6 :: perm \<Rightarrow> trm6 \<Rightarrow> trm6" | 
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 831 | is | 
| 1104 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 832 | "permute :: perm \<Rightarrow> rtrm6 \<Rightarrow> rtrm6" | 
| 
84d666f9face
the specifications of the respects.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1103diff
changeset | 833 | |
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 834 | instance | 
| 1106 | 835 | apply default | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 836 | sorry | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 837 | end | 
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 838 | |
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 839 | lemma lifted_induct: | 
| 1106 | 840 | "\<lbrakk>x1 = x2; \<And>a b. a = b \<Longrightarrow> P (Vr6 a) (Vr6 b); | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 841 | \<And>a t b s. | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 842 |    \<exists>pi. fv_trm6 t - {atom a} = fv_trm6 s - {atom b} \<and>
 | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 843 |         (fv_trm6 t - {atom a}) \<sharp>* pi \<and> pi \<bullet> t = s \<and> P (pi \<bullet> t) s \<Longrightarrow>
 | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 844 | P (Lm6 a t) (Lm6 b s); | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 845 | \<And>t1 s1 t2 s2. | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 846 | \<exists>pi. fv_trm6 s1 - bv6 t1 = fv_trm6 s2 - bv6 t2 \<and> | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 847 | (fv_trm6 s1 - bv6 t1) \<sharp>* pi \<and> pi \<bullet> s1 = s2 \<and> P (pi \<bullet> s1) s2 \<Longrightarrow> | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 848 | P (Lt6 t1 s1) (Lt6 t2 s2)\<rbrakk> | 
| 1106 | 849 | \<Longrightarrow> P x1 x2" | 
| 850 | unfolding alpha_gen | |
| 851 | apply (lifting alpha6.induct[unfolded alpha_gen]) | |
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 852 | apply injection | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 853 | (* notice unsolvable goals: (alpha6 ===> op =) rbv6 rbv6 *) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 854 | oops | 
| 1106 | 855 | |
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 856 | lemma lifted_inject_a3: | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 857 | "\<exists>pi. fv_trm6 s1 - bv6 t1 = fv_trm6 s2 - bv6 t2 \<and> | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 858 | (fv_trm6 s1 - bv6 t1) \<sharp>* pi \<and> pi \<bullet> s1 = s2 \<Longrightarrow> Lt6 t1 s1 = Lt6 t2 s2" | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 859 | apply(lifting a3[unfolded alpha_gen]) | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 860 | apply injection | 
| 1111 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 861 | (* notice unsolvable goals: (alpha6 ===> op =) rbv6 rbv6 *) | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 862 | oops | 
| 
ee276c9f12f0
A concrete example, with a proof that rbv is not regular and
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1109diff
changeset | 863 | |
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 864 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 865 | |
| 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 866 | |
| 1121 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 867 | (* example with a respectful bn function defined over the type itself *) | 
| 1131 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 868 | |
| 1121 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 869 | datatype rtrm7 = | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 870 | rVr7 "name" | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 871 | | rLm7 "name" "rtrm7" | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 872 | | rLt7 "rtrm7" "rtrm7" --"bind (bv7 left) in (right)" | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 873 | |
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 874 | primrec | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 875 | rbv7 | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 876 | where | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 877 |   "rbv7 (rVr7 n) = {atom n}"
 | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 878 | | "rbv7 (rLm7 n t) = rbv7 t - {atom n}"
 | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 879 | | "rbv7 (rLt7 l r) = rbv7 l \<union> rbv7 r" | 
| 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 880 | |
| 1180 | 881 | local_setup {* define_raw_fv "Terms.rtrm7" [
 | 
| 882 |   [[[]], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv7}, 0)]]]] *}
 | |
| 883 | print_theorems | |
| 1131 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 884 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 885 | setup {* snd o define_raw_perms ["rtrm7"] ["Terms.rtrm7"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 886 | print_theorems | 
| 1131 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 887 | |
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 888 | inductive | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 889 |   alpha7 :: "rtrm7 \<Rightarrow> rtrm7 \<Rightarrow> bool" ("_ \<approx>7 _" [100, 100] 100)
 | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 890 | where | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 891 | a1: "a = b \<Longrightarrow> (rVr7 a) \<approx>7 (rVr7 b)" | 
| 1180 | 892 | | a2: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha7 fv_rtrm7 pi ({atom b}, s))) \<Longrightarrow> rLm7 a t \<approx>7 rLm7 b s"
 | 
| 893 | | a3: "(\<exists>pi. (((rbv7 t1), s1) \<approx>gen alpha7 fv_rtrm7 pi ((rbv7 t2), s2))) \<Longrightarrow> rLt7 t1 s1 \<approx>7 rLt7 t2 s2" | |
| 1121 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 894 | |
| 1180 | 895 | lemma bvfv7: "rbv7 x = fv_rtrm7 x" | 
| 1131 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 896 | apply induct | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 897 | apply simp_all | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 898 | done | 
| 1121 
8d3f92694e85
example with a respectful bn function defined over the type itself
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1119diff
changeset | 899 | |
| 1131 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 900 | lemma "(x::name) \<noteq> y \<Longrightarrow> \<not> (alpha7 ===> op =) rbv7 rbv7" | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 901 | apply simp | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 902 | apply (rule_tac x="rLt7 (rVr7 x) (rVr7 x)" in exI) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 903 | apply (rule_tac x="rLt7 (rVr7 y) (rVr7 y)" in exI) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 904 | apply simp | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 905 | apply (rule a3) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 906 | apply (rule_tac x="(x \<leftrightarrow> y)" in exI) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 907 | apply (simp_all add: alpha_gen fresh_star_def) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 908 | apply (rule a1) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 909 | apply (rule refl) | 
| 
95e587907728
Even when bv = fv it still doesn't lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1129diff
changeset | 910 | done | 
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 911 | |
| 1133 
649680775e93
fv_foo is not regular.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1132diff
changeset | 912 | |
| 
649680775e93
fv_foo is not regular.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1132diff
changeset | 913 | |
| 
649680775e93
fv_foo is not regular.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1132diff
changeset | 914 | |
| 
649680775e93
fv_foo is not regular.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1132diff
changeset | 915 | |
| 1132 | 916 | datatype rfoo8 = | 
| 917 | Foo0 "name" | |
| 918 | | Foo1 "rbar8" "rfoo8" --"bind bv(bar) in foo" | |
| 919 | and rbar8 = | |
| 1135 | 920 | Bar0 "name" | 
| 921 | | Bar1 "name" "name" "rbar8" --"bind second name in b" | |
| 1132 | 922 | |
| 923 | primrec | |
| 924 | rbv8 | |
| 925 | where | |
| 926 |   "rbv8 (Bar0 x) = {}"
 | |
| 927 | | "rbv8 (Bar1 v x b) = {atom v}"
 | |
| 928 | ||
| 1180 | 929 | local_setup {* define_raw_fv "Terms.rfoo8" [
 | 
| 930 |   [[[]], [[], [(SOME @{term rbv8}, 0)]]], [[[]], [[], [(NONE, 1)], [(NONE, 1)]]]] *}
 | |
| 931 | print_theorems | |
| 1132 | 932 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 933 | setup {* snd o define_raw_perms ["rfoo8", "rbar8"] ["Terms.rfoo8", "Terms.rbar8"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 934 | print_theorems | 
| 1132 | 935 | |
| 936 | inductive | |
| 937 |   alpha8f :: "rfoo8 \<Rightarrow> rfoo8 \<Rightarrow> bool" ("_ \<approx>f _" [100, 100] 100)
 | |
| 938 | and | |
| 939 |   alpha8b :: "rbar8 \<Rightarrow> rbar8 \<Rightarrow> bool" ("_ \<approx>b _" [100, 100] 100)
 | |
| 940 | where | |
| 941 | a1: "a = b \<Longrightarrow> (Foo0 a) \<approx>f (Foo0 b)" | |
| 942 | | a2: "a = b \<Longrightarrow> (Bar0 a) \<approx>b (Bar0 b)" | |
| 1180 | 943 | | a3: "b1 \<approx>b b2 \<Longrightarrow> (\<exists>pi. (((rbv8 b1), t1) \<approx>gen alpha8f fv_rfoo8 pi ((rbv8 b2), t2))) \<Longrightarrow> Foo1 b1 t1 \<approx>f Foo1 b2 t2" | 
| 944 | | a4: "v1 = v2 \<Longrightarrow> (\<exists>pi. (({atom x1}, t1) \<approx>gen alpha8b fv_rbar8 pi ({atom x2}, t2))) \<Longrightarrow> Bar1 v1 x1 t1 \<approx>b Bar1 v2 x2 t2"
 | |
| 1132 | 945 | |
| 946 | lemma "(alpha8b ===> op =) rbv8 rbv8" | |
| 947 | apply simp apply clarify | |
| 948 | apply (erule alpha8f_alpha8b.inducts(2)) | |
| 949 | apply (simp_all) | |
| 950 | done | |
| 1134 
998d6b491003
Finished a working foo/bar.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1133diff
changeset | 951 | |
| 1180 | 952 | lemma fv_rbar8_rsp_hlp: "x \<approx>b y \<Longrightarrow> fv_rbar8 x = fv_rbar8 y" | 
| 1132 | 953 | apply (erule alpha8f_alpha8b.inducts(2)) | 
| 1133 
649680775e93
fv_foo is not regular.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1132diff
changeset | 954 | apply (simp_all add: alpha_gen) | 
| 1132 | 955 | done | 
| 1180 | 956 | lemma "(alpha8b ===> op =) fv_rbar8 fv_rbar8" | 
| 957 | apply simp apply clarify apply (simp add: fv_rbar8_rsp_hlp) | |
| 1134 
998d6b491003
Finished a working foo/bar.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1133diff
changeset | 958 | done | 
| 1132 | 959 | |
| 1180 | 960 | lemma "(alpha8f ===> op =) fv_rfoo8 fv_rfoo8" | 
| 1134 
998d6b491003
Finished a working foo/bar.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1133diff
changeset | 961 | apply simp apply clarify | 
| 
998d6b491003
Finished a working foo/bar.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1133diff
changeset | 962 | apply (erule alpha8f_alpha8b.inducts(1)) | 
| 1180 | 963 | apply (simp_all add: alpha_gen fv_rbar8_rsp_hlp) | 
| 1133 
649680775e93
fv_foo is not regular.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1132diff
changeset | 964 | done | 
| 1132 | 965 | |
| 966 | ||
| 967 | ||
| 1103 
7e691b3c2414
trm6 with the  'Foo' constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1093diff
changeset | 968 | |
| 1134 
998d6b491003
Finished a working foo/bar.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1133diff
changeset | 969 | |
| 1135 | 970 | |
| 971 | datatype rlam9 = | |
| 972 | Var9 "name" | |
| 973 | | Lam9 "name" "rlam9" --"bind name in rlam" | |
| 974 | and rbla9 = | |
| 975 | Bla9 "rlam9" "rlam9" --"bind bv(first) in second" | |
| 976 | ||
| 977 | primrec | |
| 978 | rbv9 | |
| 979 | where | |
| 980 |   "rbv9 (Var9 x) = {}"
 | |
| 981 | | "rbv9 (Lam9 x b) = {atom x}"
 | |
| 982 | ||
| 1180 | 983 | local_setup {* define_raw_fv "Terms.rlam9" [
 | 
| 984 |   [[[]], [[(NONE, 0)], [(NONE, 0)]]], [[[], [(SOME @{term rbv9}, 0)]]]] *}
 | |
| 985 | print_theorems | |
| 1135 | 986 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 987 | setup {* snd o define_raw_perms ["rlam9", "rbla9"] ["Terms.rlam9", "Terms.rbla9"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 988 | print_theorems | 
| 1135 | 989 | |
| 990 | inductive | |
| 991 |   alpha9l :: "rlam9 \<Rightarrow> rlam9 \<Rightarrow> bool" ("_ \<approx>9l _" [100, 100] 100)
 | |
| 992 | and | |
| 993 |   alpha9b :: "rbla9 \<Rightarrow> rbla9 \<Rightarrow> bool" ("_ \<approx>9b _" [100, 100] 100)
 | |
| 994 | where | |
| 995 | a1: "a = b \<Longrightarrow> (Var9 a) \<approx>9l (Var9 b)" | |
| 1180 | 996 | | a4: "(\<exists>pi. (({atom x1}, t1) \<approx>gen alpha9l fv_rlam9 pi ({atom x2}, t2))) \<Longrightarrow> Lam9 x1 t1 \<approx>9l Lam9 x2 t2"
 | 
| 997 | | a3: "b1 \<approx>9l b2 \<Longrightarrow> (\<exists>pi. (((rbv9 b1), t1) \<approx>gen alpha9l fv_rlam9 pi ((rbv9 b2), t2))) \<Longrightarrow> Bla9 b1 t1 \<approx>9b Bla9 b2 t2" | |
| 1135 | 998 | |
| 999 | quotient_type | |
| 1000 | lam9 = rlam9 / alpha9l and bla9 = rbla9 / alpha9b | |
| 1001 | sorry | |
| 1002 | ||
| 1003 | quotient_definition | |
| 1004 | "qVar9 :: name \<Rightarrow> lam9" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1005 | is | 
| 1135 | 1006 | "Var9" | 
| 1007 | ||
| 1008 | quotient_definition | |
| 1009 | "qLam :: name \<Rightarrow> lam9 \<Rightarrow> lam9" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1010 | is | 
| 1135 | 1011 | "Lam9" | 
| 1012 | ||
| 1013 | quotient_definition | |
| 1014 | "qBla9 :: lam9 \<Rightarrow> lam9 \<Rightarrow> bla9" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1015 | is | 
| 1135 | 1016 | "Bla9" | 
| 1017 | ||
| 1018 | quotient_definition | |
| 1019 | "fv_lam9 :: lam9 \<Rightarrow> atom set" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1020 | is | 
| 1180 | 1021 | "fv_rlam9" | 
| 1135 | 1022 | |
| 1023 | quotient_definition | |
| 1024 | "fv_bla9 :: bla9 \<Rightarrow> atom set" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1025 | is | 
| 1180 | 1026 | "fv_rbla9" | 
| 1135 | 1027 | |
| 1028 | quotient_definition | |
| 1029 | "bv9 :: lam9 \<Rightarrow> atom set" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1030 | is | 
| 1135 | 1031 | "rbv9" | 
| 1032 | ||
| 1033 | instantiation lam9 and bla9 :: pt | |
| 1034 | begin | |
| 1035 | ||
| 1036 | quotient_definition | |
| 1037 | "permute_lam9 :: perm \<Rightarrow> lam9 \<Rightarrow> lam9" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1038 | is | 
| 1135 | 1039 | "permute :: perm \<Rightarrow> rlam9 \<Rightarrow> rlam9" | 
| 1040 | ||
| 1041 | quotient_definition | |
| 1042 | "permute_bla9 :: perm \<Rightarrow> bla9 \<Rightarrow> bla9" | |
| 1139 
c4001cda9da3
renamed 'as' to 'is' everywhere.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1135diff
changeset | 1043 | is | 
| 1135 | 1044 | "permute :: perm \<Rightarrow> rbla9 \<Rightarrow> rbla9" | 
| 1045 | ||
| 1046 | instance | |
| 1047 | sorry | |
| 1048 | ||
| 1049 | end | |
| 1050 | ||
| 1051 | lemma "\<lbrakk>b1 = b2; \<exists>pi. fv_lam9 t1 - bv9 b1 = fv_lam9 t2 - bv9 b2 \<and> (fv_lam9 t1 - bv9 b1) \<sharp>* pi \<and> pi \<bullet> t1 = t2\<rbrakk> | |
| 1052 | \<Longrightarrow> qBla9 b1 t1 = qBla9 b2 t2" | |
| 1053 | apply (lifting a3[unfolded alpha_gen]) | |
| 1054 | apply injection | |
| 1055 | sorry | |
| 1056 | ||
| 1057 | ||
| 1058 | ||
| 1059 | ||
| 1060 | ||
| 1061 | ||
| 1062 | ||
| 1063 | ||
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1064 | text {* type schemes *} 
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1065 | datatype ty = | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1066 | Var "name" | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1067 | | Fun "ty" "ty" | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1068 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 1069 | setup {* snd o define_raw_perms ["ty"] ["Terms.ty"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 1070 | print_theorems | 
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1071 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1072 | datatype tyS = | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1073 | All "name set" "ty" | 
| 1046 | 1074 | |
| 1171 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 1075 | setup {* snd o define_raw_perms ["tyS"] ["Terms.tyS"] *}
 | 
| 
62632eec979c
Tested the Perm code; works everywhere in Terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1139diff
changeset | 1076 | print_theorems | 
| 1040 
632467a97dd8
alpha5 pseudo-injective
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1036diff
changeset | 1077 | |
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1078 | abbreviation | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1079 |   "atoms xs \<equiv> {atom x| x. x \<in> xs}"
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1080 | |
| 1180 | 1081 | local_setup {* define_raw_fv "Terms.ty" [[[[]], [[], []]]] *}
 | 
| 1082 | print_theorems | |
| 1083 | ||
| 1084 | (* | |
| 1085 | doesn't work yet | |
| 1086 | local_setup {* define_raw_fv "Terms.tyS" [[[[], []]]] *}
 | |
| 1087 | print_theorems | |
| 1088 | *) | |
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1089 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1090 | primrec | 
| 1180 | 1091 | fv_tyS | 
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1092 | where | 
| 1180 | 1093 | "fv_tyS (All xs T) = (fv_ty T - atoms xs)" | 
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1094 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1095 | inductive | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1096 |   alpha_tyS :: "tyS \<Rightarrow> tyS \<Rightarrow> bool" ("_ \<approx>tyS _" [100, 100] 100)
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1097 | where | 
| 1180 | 1098 | a1: "\<exists>pi. ((atoms xs1, T1) \<approx>gen (op =) fv_ty pi (atoms xs2, T2)) | 
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1099 | \<Longrightarrow> All xs1 T1 \<approx>tyS All xs2 T2" | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1100 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1101 | lemma | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1102 |   shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {b, a} (Fun (Var a) (Var b))"
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1103 | apply(rule a1) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1104 | apply(simp add: alpha_gen) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1105 | apply(rule_tac x="0::perm" in exI) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1106 | apply(simp add: fresh_star_def) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1107 | done | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1108 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1109 | lemma | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1110 |   shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var b) (Var a))"
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1111 | apply(rule a1) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1112 | apply(simp add: alpha_gen) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1113 | apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1114 | apply(simp add: fresh_star_def) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1115 | done | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1116 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1117 | lemma | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1118 |   shows "All {a, b, c} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var a) (Var b))"
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1119 | apply(rule a1) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1120 | apply(simp add: alpha_gen) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1121 | apply(rule_tac x="0::perm" in exI) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1122 | apply(simp add: fresh_star_def) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1123 | done | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1124 | |
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1125 | lemma | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1126 | assumes a: "a \<noteq> b" | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1127 |   shows "\<not>(All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {c} (Fun (Var c) (Var c)))"
 | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1128 | using a | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1129 | apply(clarify) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1130 | apply(erule alpha_tyS.cases) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1131 | apply(simp add: alpha_gen) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1132 | apply(erule conjE)+ | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1133 | apply(erule exE) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1134 | apply(erule conjE)+ | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1135 | apply(clarify) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1136 | apply(simp) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1137 | apply(simp add: fresh_star_def) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1138 | apply(auto) | 
| 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1139 | done | 
| 1051 | 1140 | |
| 1048 
f5e037fd7c01
added type-scheme example
 Christian Urban <urbanc@in.tum.de> parents: 
1044diff
changeset | 1141 | |
| 1035 
3a60a028cfc5
Starting with a let-rec example.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1033diff
changeset | 1142 | end |