| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Tue, 02 Mar 2010 11:04:49 +0100 | |
| changeset 1300 | 22a084c9316b | 
| parent 1291 | 24889782da92 | 
| child 1343 | 693df83172f0 | 
| permissions | -rw-r--r-- | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 1 | theory Term1 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 2 | imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 5 | atom_decl name | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 6 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | section {*** lets with binding patterns ***}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 9 | datatype rtrm1 = | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 10 | rVr1 "name" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 11 | | rAp1 "rtrm1" "rtrm1" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 12 | | rLm1 "name" "rtrm1" --"name is bound in trm1" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 13 | | rLt1 "bp" "rtrm1" "rtrm1" --"all variables in bp are bound in the 2nd trm1" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 14 | and bp = | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 15 | BUnit | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 16 | | BVr "name" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 17 | | BPr "bp" "bp" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 18 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 19 | print_theorems | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 20 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 21 | (* to be given by the user *) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 22 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 23 | primrec | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 24 | bv1 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 25 | where | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 26 |   "bv1 (BUnit) = {}"
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 27 | | "bv1 (BVr x) = {atom x}"
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 28 | | "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 29 | |
| 1277 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1274diff
changeset | 30 | setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term1.rtrm1") 2 *}
 | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 31 | thm permute_rtrm1_permute_bp.simps | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 32 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 33 | local_setup {*
 | 
| 1277 
6eacf60ce41d
Permutation and FV_Alpha interface change.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1274diff
changeset | 34 |   snd o define_fv_alpha (Datatype.the_info @{theory} "Term1.rtrm1")
 | 
| 1291 
24889782da92
Trying to prove equivariance.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1283diff
changeset | 35 |   [[[], [], [(NONE, 0, 1)], [(SOME @{term bv1}, 0, 2)]],
 | 
| 
24889782da92
Trying to prove equivariance.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1283diff
changeset | 36 | [[], [], []]] *} | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 37 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 38 | notation | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 |   alpha_rtrm1 ("_ \<approx>1 _" [100, 100] 100) and
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 40 |   alpha_bp ("_ \<approx>1b _" [100, 100] 100)
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 41 | thm alpha_rtrm1_alpha_bp.intros | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1278diff
changeset | 42 | thm fv_rtrm1_fv_bp.simps[no_vars] | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 43 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 44 | local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_inj}, []), (build_alpha_inj @{thms alpha_rtrm1_alpha_bp.intros} @{thms rtrm1.distinct rtrm1.inject bp.distinct bp.inject} @{thms alpha_rtrm1.cases alpha_bp.cases} ctxt)) ctxt)) *}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 45 | thm alpha1_inj | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 46 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 47 | local_setup {*
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 48 | snd o (build_eqvts @{binding bv1_eqvt} [@{term bv1}] [@{term "permute :: perm \<Rightarrow> bp \<Rightarrow> bp"}] (@{thms bv1.simps permute_rtrm1_permute_bp.simps}) @{thm rtrm1_bp.inducts(2)})
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 49 | *} | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | local_setup {*
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 52 | snd o build_eqvts @{binding fv_rtrm1_fv_bp_eqvt} [@{term fv_rtrm1}, @{term fv_bp}] [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"},@{term "permute :: perm \<Rightarrow> bp \<Rightarrow> bp"}] (@{thms fv_rtrm1_fv_bp.simps permute_rtrm1_permute_bp.simps}) @{thm rtrm1_bp.induct}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 53 | *} | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 54 | |
| 1300 | 55 | local_setup {*
 | 
| 56 | (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_eqvt}, []),
 | |
| 57 | build_alpha_eqvts [@{term alpha_rtrm1}, @{term alpha_bp}] [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"},@{term "permute :: perm \<Rightarrow> bp \<Rightarrow> bp"}] @{thms permute_rtrm1_permute_bp.simps alpha1_inj} @{thm alpha_rtrm1_alpha_bp.induct} ctxt) ctxt)) *}
 | |
| 1291 
24889782da92
Trying to prove equivariance.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1283diff
changeset | 58 | |
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 59 | lemma alpha1_eqvt_proper[eqvt]: | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 60 | "pi \<bullet> (t \<approx>1 s) = ((pi \<bullet> t) \<approx>1 (pi \<bullet> s))" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 61 | "pi \<bullet> (alpha_bp a b) = (alpha_bp (pi \<bullet> a) (pi \<bullet> b))" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 62 | apply (simp_all only: eqvts) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 63 | apply rule | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 64 | apply (simp_all add: alpha1_eqvt) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 65 | apply (subst permute_minus_cancel(2)[symmetric,of "t" "pi"]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 66 | apply (subst permute_minus_cancel(2)[symmetric,of "s" "pi"]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 67 | apply (simp_all only: alpha1_eqvt) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 68 | apply rule | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 69 | apply (simp_all add: alpha1_eqvt) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 70 | apply (subst permute_minus_cancel(2)[symmetric,of "a" "pi"]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 71 | apply (subst permute_minus_cancel(2)[symmetric,of "b" "pi"]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 72 | apply (simp_all only: alpha1_eqvt) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | done | 
| 1274 
d867021d8ac1
Preparing the generalized lifting procedure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1270diff
changeset | 74 | thm eqvts_raw(1) | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 76 | local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_equivp}, []),
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 77 |   (build_equivps [@{term alpha_rtrm1}, @{term alpha_bp}] @{thm rtrm1_bp.induct} @{thm alpha_rtrm1_alpha_bp.induct} @{thms rtrm1.inject bp.inject} @{thms alpha1_inj} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} ctxt)) ctxt)) *}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 78 | thm alpha1_equivp | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 | local_setup  {* define_quotient_type [(([], @{binding trm1}, NoSyn), (@{typ rtrm1}, @{term alpha_rtrm1}))]
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 |   (rtac @{thm alpha1_equivp(1)} 1) *}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 82 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 83 | local_setup {*
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 84 | (fn ctxt => ctxt | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 85 |  |> snd o (Quotient_Def.quotient_lift_const ("Vr1", @{term rVr1}))
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 86 |  |> snd o (Quotient_Def.quotient_lift_const ("Ap1", @{term rAp1}))
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 87 |  |> snd o (Quotient_Def.quotient_lift_const ("Lm1", @{term rLm1}))
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 88 |  |> snd o (Quotient_Def.quotient_lift_const ("Lt1", @{term rLt1}))
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 89 |  |> snd o (Quotient_Def.quotient_lift_const ("fv_trm1", @{term fv_rtrm1})))
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | *} | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | print_theorems | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 92 | |
| 1278 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 93 | local_setup {* snd o prove_const_rsp @{binding fv_rtrm1_rsp} [@{term fv_rtrm1}]
 | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 94 |   (fn _ => fvbv_rsp_tac @{thm alpha_rtrm1_alpha_bp.inducts(1)} @{thms fv_rtrm1_fv_bp.simps} 1) *}
 | 
| 1278 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 95 | local_setup {* snd o prove_const_rsp @{binding rVr1_rsp} [@{term rVr1}]
 | 
| 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 96 |   (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
 | 
| 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 97 | local_setup {* snd o prove_const_rsp @{binding rAp1_rsp} [@{term rAp1}]
 | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 98 |   (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
 | 
| 1278 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 99 | local_setup {* snd o prove_const_rsp @{binding rLm1_rsp} [@{term rLm1}]
 | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 100 |   (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
 | 
| 1278 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 101 | local_setup {* snd o prove_const_rsp @{binding rLt1_rsp} [@{term rLt1}]
 | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 102 |   (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
 | 
| 1278 
8814494fe4da
Change in signature of prove_const_rsp for general lifting.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 103 | local_setup {* snd o prove_const_rsp @{binding permute_rtrm1_rsp} [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"}]
 | 
| 1270 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 104 |   (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha1_eqvt}) 1) *}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 105 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 106 | lemmas trm1_bp_induct = rtrm1_bp.induct[quot_lifted] | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 107 | lemmas trm1_bp_inducts = rtrm1_bp.inducts[quot_lifted] | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 108 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 109 | setup {* define_lifted_perms ["Term1.trm1"] [("permute_trm1", @{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"})]
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 110 |   @{thms permute_rtrm1_permute_bp_zero permute_rtrm1_permute_bp_append} *}
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 111 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 112 | lemmas | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 113 | permute_trm1 = permute_rtrm1_permute_bp.simps[quot_lifted] | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 114 | and fv_trm1 = fv_rtrm1_fv_bp.simps[quot_lifted] | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 115 | and fv_trm1_eqvt = fv_rtrm1_fv_bp_eqvt(1)[quot_lifted] | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 116 | and alpha1_INJ = alpha1_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 117 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 118 | lemma supports: | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 119 | "(supp (atom x)) supports (Vr1 x)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 120 | "(supp t \<union> supp s) supports (Ap1 t s)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 121 | "(supp (atom x) \<union> supp t) supports (Lm1 x t)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 122 | "(supp b \<union> supp t \<union> supp s) supports (Lt1 b t s)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 123 |   "{} supports BUnit"
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 124 | "(supp (atom x)) supports (BVr x)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 125 | "(supp a \<union> supp b) supports (BPr a b)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 126 | apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh permute_trm1) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 127 | apply(rule_tac [!] allI)+ | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 128 | apply(rule_tac [!] impI) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 129 | apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | apply(simp_all add: fresh_atom) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 132 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 133 | lemma rtrm1_bp_fs: | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 134 | "finite (supp (x :: trm1))" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 135 | "finite (supp (b :: bp))" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 136 | apply (induct x and b rule: trm1_bp_inducts) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 137 |   apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 138 | apply(simp_all add: supp_atom) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 139 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 140 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 141 | instance trm1 :: fs | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 142 | apply default | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 143 | apply (rule rtrm1_bp_fs(1)) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 144 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 145 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 146 | lemma fv_eq_bv: "fv_bp bp = bv1 bp" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 147 | apply(induct bp rule: trm1_bp_inducts(2)) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 148 | apply(simp_all) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 149 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 150 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 151 | lemma helper2: "{b. \<forall>pi. pi \<bullet> (a \<rightleftharpoons> b) \<bullet> bp \<noteq> bp} = {}"
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 152 | apply auto | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 153 | apply (rule_tac x="(x \<rightleftharpoons> a)" in exI) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 154 | apply auto | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 155 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 156 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 157 | lemma alpha_bp_eq_eq: "alpha_bp a b = (a = b)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 158 | apply rule | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 159 | apply (induct a b rule: alpha_rtrm1_alpha_bp.inducts(2)) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 160 | apply (simp_all add: equivp_reflp[OF alpha1_equivp(2)]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 161 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 162 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 163 | lemma alpha_bp_eq: "alpha_bp = (op =)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 164 | apply (rule ext)+ | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 165 | apply (rule alpha_bp_eq_eq) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 166 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 167 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 168 | lemma supp_fv: | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 169 | "supp t = fv_trm1 t" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 170 | "supp b = fv_bp b" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 171 | apply(induct t and b rule: trm1_bp_inducts) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 172 | apply(simp_all) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 173 | apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 174 | apply(simp only: supp_at_base[simplified supp_def]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 175 | apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 176 | apply(simp add: Collect_imp_eq Collect_neg_eq) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 177 | apply(subgoal_tac "supp (Lm1 name rtrm1) = supp (Abs {atom name} rtrm1)")
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 178 | apply(simp add: supp_Abs fv_trm1) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 179 | apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt permute_trm1) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 180 | apply(simp add: alpha1_INJ) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 181 | apply(simp add: Abs_eq_iff) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 182 | apply(simp add: alpha_gen.simps) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 183 | apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 184 | apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp(rtrm11) \<union> supp (Abs (bv1 bp) rtrm12)") | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 185 | apply(simp add: supp_Abs fv_trm1 fv_eq_bv) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 186 | apply(simp (no_asm) add: supp_def permute_trm1) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 187 | apply(simp add: alpha1_INJ alpha_bp_eq) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 188 | apply(simp add: Abs_eq_iff) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 189 | apply(simp add: alpha_gen) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 190 | apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 191 | apply(simp add: Collect_imp_eq Collect_neg_eq fresh_star_def helper2) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 192 | apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 193 | apply(simp (no_asm) add: supp_def eqvts) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 194 | apply(fold supp_def) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 195 | apply(simp add: supp_at_base) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 196 | apply(simp (no_asm) add: supp_def Collect_imp_eq Collect_neg_eq) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 197 | apply(simp add: Collect_imp_eq[symmetric] Collect_neg_eq[symmetric] supp_def[symmetric]) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 198 | done | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 199 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 200 | lemma trm1_supp: | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 201 |   "supp (Vr1 x) = {atom x}"
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 202 | "supp (Ap1 t1 t2) = supp t1 \<union> supp t2" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 203 |   "supp (Lm1 x t) = (supp t) - {atom x}"
 | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 204 | "supp (Lt1 b t s) = supp t \<union> (supp s - bv1 b)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 205 | by (simp_all add: supp_fv fv_trm1 fv_eq_bv) | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 206 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 207 | lemma trm1_induct_strong: | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 208 | assumes "\<And>name b. P b (Vr1 name)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 209 | and "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 210 | and "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 211 | 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)" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 212 | shows "P a rtrma" | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 213 | sorry | 
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 214 | |
| 
8c3cf9f4f5f2
Split Terms into separate files and add them to tests.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 215 | end |