818 also have "\<dots> = Es2\<down> \<odot> (Es1\<down> \<odot> E)" using odot_assoc by simp |
818 also have "\<dots> = Es2\<down> \<odot> (Es1\<down> \<odot> E)" using odot_assoc by simp |
819 also have "\<dots> = Es2\<down> \<odot> (E # Es1)\<down>" by simp |
819 also have "\<dots> = Es2\<down> \<odot> (E # Es1)\<down>" by simp |
820 finally show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" by simp |
820 finally show "((E # Es1) @ Es2)\<down> = Es2\<down> \<odot> (E # Es1)\<down>" by simp |
821 qed |
821 qed |
822 |
822 |
823 text {****************************************************************** |
|
824 |
|
825 Formalising Barendregt's Proof of the Substitution Lemma |
|
826 -------------------------------------------------------- |
|
827 |
|
828 Barendregt's proof needs in the variable case a case distinction. |
|
829 One way to do this in Isar is to use blocks. A block is some sequent |
|
830 or reasoning steps enclosed in curly braces |
|
831 |
|
832 { \<dots> |
|
833 |
|
834 have "statement" |
|
835 } |
|
836 |
|
837 Such a block can contain local assumptions like |
|
838 |
|
839 { assume "A" |
|
840 assume "B" |
|
841 \<dots> |
|
842 have "C" by \<dots> |
|
843 } |
|
844 |
|
845 Where "C" is the last have-statement in this block. The behaviour |
|
846 of such a block to the 'outside' is the implication |
|
847 |
|
848 \<lbrakk>A; B\<rbrakk> \<Longrightarrow> "C" |
|
849 |
|
850 Now if we want to prove a property "smth" using the case-distinctions |
|
851 P1, P2 and P3 then we can use the following reasoning: |
|
852 |
|
853 { assume "P1" |
|
854 \<dots> |
|
855 have "smth" |
|
856 } |
|
857 moreover |
|
858 { assume "P2" |
|
859 \<dots> |
|
860 have "smth" |
|
861 } |
|
862 moreover |
|
863 { assume "P3" |
|
864 \<dots> |
|
865 have "smth" |
|
866 } |
|
867 ultimately have "smth" by blast |
|
868 |
|
869 The blocks establish the implications |
|
870 |
|
871 P1 \<Longrightarrow> smth |
|
872 P2 \<Longrightarrow> smth |
|
873 P3 \<Longrightarrow> smth |
|
874 |
|
875 If we know that P1, P2 and P3 cover all the cases, that is P1 \<or> P2 \<or> P3 is |
|
876 true, then we have 'ultimately' established the property "smth" |
|
877 |
|
878 *} |
|
879 |
|
880 section {* EXERCISE 7 *} |
|
881 |
|
882 text {* |
|
883 Fill in the cases 1.2 and 1.3 and the equational reasoning |
|
884 in the lambda-case. |
|
885 *} |
|
886 |
|
887 lemma forget: |
|
888 shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t" |
|
889 by (nominal_induct t avoiding: x s rule: lam.strong_induct) |
|
890 (auto simp add: lam.fresh fresh_at_base) |
|
891 |
|
892 lemma fresh_fact: |
|
893 assumes a: "atom z \<sharp> s" |
|
894 and b: "z = y \<or> atom z \<sharp> t" |
|
895 shows "atom z \<sharp> t[y ::= s]" |
|
896 using a b |
|
897 by (nominal_induct t avoiding: z y s rule: lam.strong_induct) |
|
898 (auto simp add: lam.fresh fresh_at_base) |
|
899 |
|
900 |
|
901 lemma |
|
902 assumes a: "x \<noteq> y" |
|
903 and b: "atom x \<sharp> L" |
|
904 shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
905 using a b |
|
906 proof (nominal_induct M avoiding: x y N L rule: lam.strong_induct) |
|
907 case (Var z) |
|
908 have a1: "x \<noteq> y" by fact |
|
909 have a2: "atom x \<sharp> L" by fact |
|
910 show "Var z[x::=N][y::=L] = Var z[y::=L][x::=N[y::=L]]" (is "?LHS = ?RHS") |
|
911 proof - |
|
912 { -- {* Case 1.1 *} |
|
913 assume c1: "z = x" |
|
914 have "(1)": "?LHS = N[y::=L]" using c1 by simp |
|
915 have "(2)": "?RHS = N[y::=L]" using c1 a1 by simp |
|
916 have "?LHS = ?RHS" using "(1)" "(2)" by simp |
|
917 } |
|
918 moreover |
|
919 { -- {* Case 1.2 *} |
|
920 assume c2: "z = y" "z \<noteq> x" |
|
921 |
|
922 have "?LHS = ?RHS" sorry |
|
923 } |
|
924 moreover |
|
925 { -- {* Case 1.3 *} |
|
926 assume c3: "z \<noteq> x" "z \<noteq> y" |
|
927 |
|
928 have "?LHS = ?RHS" sorry |
|
929 } |
|
930 ultimately show "?LHS = ?RHS" by blast |
|
931 qed |
|
932 next |
|
933 case (Lam z M1) -- {* case 2: lambdas *} |
|
934 have ih: "\<lbrakk>x \<noteq> y; atom x \<sharp> L\<rbrakk> \<Longrightarrow> M1[x::=N][y::=L] = M1[y::=L][x::=N[y::=L]]" by fact |
|
935 have a1: "x \<noteq> y" by fact |
|
936 have a2: "atom x \<sharp> L" by fact |
|
937 have fs: "atom z \<sharp> x" "atom z \<sharp> y" "atom z \<sharp> N" "atom z \<sharp> L" by fact+ |
|
938 then have b: "atom z \<sharp> N[y::=L]" by (simp add: fresh_fact) |
|
939 show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS") |
|
940 proof - |
|
941 have "?LHS = \<dots>" sorry |
|
942 |
|
943 also have "\<dots> = ?RHS" sorry |
|
944 finally show "?LHS = ?RHS" by simp |
|
945 qed |
|
946 next |
|
947 case (App M1 M2) -- {* case 3: applications *} |
|
948 then show "(App M1 M2)[x::=N][y::=L] = (App M1 M2)[y::=L][x::=N[y::=L]]" by simp |
|
949 qed |
|
950 |
|
951 text {* |
|
952 Again the strong induction principle enables Isabelle to find |
|
953 the proof of the substitution lemma automatically. |
|
954 *} |
|
955 |
|
956 lemma substitution_lemma_version: |
|
957 assumes asm: "x \<noteq> y" "atom x \<sharp> L" |
|
958 shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
959 using asm |
|
960 by (nominal_induct M avoiding: x y N L rule: lam.strong_induct) |
|
961 (auto simp add: fresh_fact forget) |
|
962 |
|
963 |
823 |
964 |
824 |
965 end |
825 end |
966 |
826 |