822 val Fake_parts_insert = @{thm Fake_parts_insert}; |
822 val Fake_parts_insert = @{thm Fake_parts_insert}; |
823 val Fake_analz_insert = @{thm Fake_analz_insert}; |
823 val Fake_analz_insert = @{thm Fake_analz_insert}; |
824 val pushes = @{thms pushes}; |
824 val pushes = @{thms pushes}; |
825 |
825 |
826 |
826 |
827 (*Prove base case (subgoal i) and simplify others. A typical base case |
|
828 concerns Crypt K X \<notin> Key`shrK`bad and cannot be proved by rewriting |
|
829 alone.*) |
|
830 fun prove_simple_subgoals_tac (cs, ss) i = |
|
831 force_tac (cs, ss addsimps [@{thm image_eq_UN}]) i THEN |
|
832 ALLGOALS (asm_simp_tac ss) |
|
833 |
|
834 (*Analysis of Fake cases. Also works for messages that forward unknown parts, |
827 (*Analysis of Fake cases. Also works for messages that forward unknown parts, |
835 but this application is no longer necessary if analz_insert_eq is used. |
828 but this application is no longer necessary if analz_insert_eq is used. |
836 Abstraction over i is ESSENTIAL: it delays the dereferencing of claset |
829 Abstraction over i is ESSENTIAL: it delays the dereferencing of claset |
837 DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) |
830 DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) |
838 |
831 |
847 eresolve_tac [asm_rl, @{thm synth.Inj}]; |
840 eresolve_tac [asm_rl, @{thm synth.Inj}]; |
848 |
841 |
849 fun Fake_insert_simp_tac ss i = |
842 fun Fake_insert_simp_tac ss i = |
850 REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i; |
843 REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i; |
851 |
844 |
852 fun atomic_spy_analz_tac (cs,ss) = SELECT_GOAL |
845 fun atomic_spy_analz_tac ctxt = |
853 (Fake_insert_simp_tac ss 1 |
846 let val ss = simpset_of ctxt and cs = claset_of ctxt in |
854 THEN |
847 SELECT_GOAL |
855 IF_UNSOLVED (Blast.depth_tac |
848 (Fake_insert_simp_tac ss 1 |
856 (cs addIs [analz_insertI, |
849 THEN |
857 impOfSubs analz_subset_parts]) 4 1)) |
850 IF_UNSOLVED (Blast.depth_tac |
858 |
851 (cs addIs [analz_insertI, |
859 fun spy_analz_tac (cs,ss) i = |
852 impOfSubs analz_subset_parts]) 4 1)) |
860 DETERM |
853 end; |
861 (SELECT_GOAL |
854 |
862 (EVERY |
855 fun spy_analz_tac ctxt i = |
863 [ (*push in occurrences of X...*) |
856 let val ss = simpset_of ctxt and cs = claset_of ctxt in |
864 (REPEAT o CHANGED) |
857 DETERM |
865 (res_inst_tac (Simplifier.the_context ss) |
858 (SELECT_GOAL |
866 [(("x", 1), "X")] (insert_commute RS ssubst) 1), |
859 (EVERY |
867 (*...allowing further simplifications*) |
860 [ (*push in occurrences of X...*) |
868 simp_tac ss 1, |
861 (REPEAT o CHANGED) |
869 REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])), |
862 (res_inst_tac ctxt [(("x", 1), "X")] (insert_commute RS ssubst) 1), |
870 DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i) |
863 (*...allowing further simplifications*) |
|
864 simp_tac ss 1, |
|
865 REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])), |
|
866 DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i) |
|
867 end; |
871 *} |
868 *} |
872 |
869 |
873 text{*By default only @{text o_apply} is built-in. But in the presence of |
870 text{*By default only @{text o_apply} is built-in. But in the presence of |
874 eta-expansion this means that some terms displayed as @{term "f o g"} will be |
871 eta-expansion this means that some terms displayed as @{term "f o g"} will be |
875 rewritten, and others will not!*} |
872 rewritten, and others will not!*} |
913 done |
910 done |
914 |
911 |
915 lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD] |
912 lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD] |
916 |
913 |
917 method_setup spy_analz = {* |
914 method_setup spy_analz = {* |
918 Scan.succeed (SIMPLE_METHOD' o spy_analz_tac o clasimpset_of) *} |
915 Scan.succeed (SIMPLE_METHOD' o spy_analz_tac) *} |
919 "for proving the Fake case when analz is involved" |
916 "for proving the Fake case when analz is involved" |
920 |
917 |
921 method_setup atomic_spy_analz = {* |
918 method_setup atomic_spy_analz = {* |
922 Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac o clasimpset_of) *} |
919 Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac) *} |
923 "for debugging spy_analz" |
920 "for debugging spy_analz" |
924 |
921 |
925 method_setup Fake_insert_simp = {* |
922 method_setup Fake_insert_simp = {* |
926 Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o simpset_of) *} |
923 Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o simpset_of) *} |
927 "for debugging spy_analz" |
924 "for debugging spy_analz" |